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,677 @@ (* 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-Complex_Analysis.Complex_Analysis" "HOL-Number_Theory.Number_Theory" "HOL-Computational_Algebra.Polynomial" begin +hide_const (open) UnivPoly.coeff UnivPoly.up_ring.monom +hide_const (open) Module.smult Coset.order + (* 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) + by (metis dvdI minus_one_mult_self unit_imp_no_prime_divisors) lemma nat_fact [simp]: "nat (fact n) = fact n" - by (subst of_nat_fact [symmetric]) (rule nat_int) + by (metis nat_int of_nat_fact of_nat_fact) 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 + by (simp add: of_int_poly.rep_eq) 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 + by (meson dvd_trans fact_dvd fact_dvd_poly_higher_pderiv_aux) 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 obtain N where "(\(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) with that show thesis by blast qed 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]) + by (intro filterlim_compose[OF power_over_fact_tendsto_0' filterlim_minus_const_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/Formal_Puiseux_Series/Formal_Puiseux_Series.thy b/thys/Formal_Puiseux_Series/Formal_Puiseux_Series.thy --- a/thys/Formal_Puiseux_Series/Formal_Puiseux_Series.thy +++ b/thys/Formal_Puiseux_Series/Formal_Puiseux_Series.thy @@ -1,1824 +1,1824 @@ (* File: Formal_Puiseux_Series.thy Author: Manuel Eberl, TU München *) section \Formal Puiseux Series\ theory Formal_Puiseux_Series - imports Puiseux_Laurent_Library FPS_Hensel + imports FPS_Hensel begin subsection \Auxiliary facts and definitions\ lemma div_dvd_self: fixes a b :: "'a :: {semidom_divide}" shows "b dvd a \ a div b dvd a" by (elim dvdE; cases "b = 0") simp_all lemma quotient_of_int [simp]: "quotient_of (of_int n) = (n, 1)" using Rat.of_int_def quotient_of_int by auto lemma of_int_div_of_int_in_Ints_iff: "(of_int n / of_int m :: 'a :: field_char_0) \ \ \ m = 0 \ m dvd n" proof assume *: "(of_int n / of_int m :: 'a) \ \" { assume "m \ 0" from * obtain k where k: "(of_int n / of_int m :: 'a) = of_int k" by (auto elim!: Ints_cases) hence "of_int n = (of_int k * of_int m :: 'a)" using \m \ 0\ by (simp add: field_simps) also have "\ = of_int (k * m)" by simp finally have "n = k * m" by (subst (asm) of_int_eq_iff) hence "m dvd n" by auto } thus "m = 0 \ m dvd n" by blast qed auto lemma rat_eq_quotientD: assumes "r = rat_of_int a / rat_of_int b" "b \ 0" shows "fst (quotient_of r) dvd a" "snd (quotient_of r) dvd b" proof - define a' b' where "a' = fst (quotient_of r)" and "b' = snd (quotient_of r)" define d where "d = gcd a b" have "b' > 0" by (auto simp: b'_def quotient_of_denom_pos') have "coprime a' b'" by (rule quotient_of_coprime[of r]) (simp add: a'_def b'_def) have r: "r = rat_of_int a' / rat_of_int b'" by (simp add: a'_def b'_def quotient_of_div) from assms \b' > 0\ have "rat_of_int (a' * b) = rat_of_int (a * b')" unfolding of_int_mult by (simp add: field_simps r) hence eq: "a' * b = a * b'" by (subst (asm) of_int_eq_iff) have "a' dvd a * b'" by (simp flip: eq) hence "a' dvd a" by (subst (asm) coprime_dvd_mult_left_iff) fact moreover have "b' dvd a' * b" by (simp add: eq) hence "b' dvd b" by (subst (asm) coprime_dvd_mult_right_iff) (use \coprime a' b'\ in \simp add: coprime_commute\) ultimately show "fst (quotient_of r) dvd a" "snd (quotient_of r) dvd b" unfolding a'_def b'_def by blast+ qed lemma quotient_of_denom_add_dvd: "snd (quotient_of (x + y)) dvd snd (quotient_of x) * snd (quotient_of y)" proof - define a b where "a = fst (quotient_of x)" and "b = snd (quotient_of x)" define c d where "c = fst (quotient_of y)" and "d = snd (quotient_of y)" have "b > 0" "d > 0" by (auto simp: b_def d_def quotient_of_denom_pos') have xy: "x = rat_of_int a / rat_of_int b" "y = rat_of_int c / rat_of_int d" unfolding a_def b_def c_def d_def by (simp_all add: quotient_of_div) show "snd (quotient_of (x + y)) dvd b * d" proof (rule rat_eq_quotientD) show "x + y = rat_of_int (a * d + c * b) / rat_of_int (b * d)" using \b > 0\ \d > 0\ by (simp add: field_simps xy) qed (use \b > 0\ \d > 0\ in auto) qed lemma quotient_of_denom_diff_dvd: "snd (quotient_of (x - y)) dvd snd (quotient_of x) * snd (quotient_of y)" using quotient_of_denom_add_dvd[of x "-y"] by (simp add: rat_uminus_code Let_def case_prod_unfold) definition supp :: "('a \ ('b :: zero)) \ 'a set" where "supp f = f -` (-{0})" lemma supp_0 [simp]: "supp (\_. 0) = {}" and supp_const: "supp (\_. c) = (if c = 0 then {} else UNIV)" and supp_singleton [simp]: "c \ 0 \ supp (\x. if x = d then c else 0) = {d}" by (auto simp: supp_def) lemma supp_uminus [simp]: "supp (\x. -f x :: 'a :: group_add) = supp f" by (auto simp: supp_def) subsection \Definition\ text \ Similarly to formal power series $R[[X]]$ and formal Laurent series $R((X))$, we define the ring of formal Puiseux series $R\{\{X\}\}$ as functions from the rationals into a ring such that \<^enum> the support is bounded from below, and \<^enum> the denominators of the numbers in the support have a common multiple other than 0 One can also think of a formal Puiseux series in the paramter $X$ as a formal Laurent series in the parameter $X^{1/d}$ for some positive integer $d$. This is often written in the following suggestive notation: \[ R\{\{X\}\} = \bigcup_{d\geq 1} R((X^{1/d})) \] Many operations will be defined in terms of this correspondence between Puiseux and Laurent series, and many of the simple properties proven that way. \ definition is_fpxs :: "(rat \ 'a :: zero) \ bool" where "is_fpxs f \ bdd_below (supp f) \ (LCM r\supp f. snd (quotient_of r)) \ 0" typedef (overloaded) 'a fpxs = "{f::rat \ 'a :: zero. is_fpxs f}" morphisms fpxs_nth Abs_fpxs by (rule exI[of _ "\_. 0"]) (auto simp: is_fpxs_def supp_def) setup_lifting type_definition_fpxs lemma fpxs_ext: "(\r. fpxs_nth f r = fpxs_nth g r) \ f = g" by transfer auto lemma fpxs_eq_iff: "f = g \ (\r. fpxs_nth f r = fpxs_nth g r)" by transfer auto lift_definition fpxs_supp :: "'a :: zero fpxs \ rat set" is supp . lemma fpxs_supp_altdef: "fpxs_supp f = {x. fpxs_nth f x \ 0}" by transfer (auto simp: supp_def) text \ The following gives us the ``root order'' of \f\i, i.e. the smallest positive integer \d\ such that \f\ is in $R((X^{1/p}))$. \ lift_definition fpxs_root_order :: "'a :: zero fpxs \ nat" is "\f. nat (LCM r\supp f. snd (quotient_of r))" . lemma fpxs_root_order_pos [simp]: "fpxs_root_order f > 0" proof transfer fix f :: "rat \ 'a" assume f: "is_fpxs f" hence "(LCM r\supp f. snd (quotient_of r)) \ 0" by (auto simp: is_fpxs_def) moreover have "(LCM r\supp f. snd (quotient_of r)) \ 0" by simp ultimately show "nat (LCM r\supp f. snd (quotient_of r)) > 0" by linarith qed lemma fpxs_root_order_nonzero [simp]: "fpxs_root_order f \ 0" using fpxs_root_order_pos[of f] by linarith text \ Let \d\ denote the root order of a Puiseux series \f\, i.e. the smallest number \d\ such that all monomials with non-zero coefficients can be written in the form $X^{n/d}$ for some \n\. Then \f\ can be written as a Laurent series in \X^{1/d}\. The following operation gives us this Laurent series. \ lift_definition fls_of_fpxs :: "'a :: zero fpxs \ 'a fls" is "\f n. f (of_int n / of_int (LCM r\supp f. snd (quotient_of r)))" proof - fix f :: "rat \ 'a" assume f: "is_fpxs f" hence "bdd_below (supp f)" by (auto simp: is_fpxs_def) then obtain r0 where "\x\supp f. r0 \ x" by (auto simp: bdd_below_def) hence r0: "f x = 0" if "x < r0" for x using that by (auto simp: supp_def vimage_def) define d :: int where "d = (LCM r\supp f. snd (quotient_of r))" have "d \ 0" by (simp add: d_def) moreover have "d \ 0" using f by (auto simp: d_def is_fpxs_def) ultimately have "d > 0" by linarith have *: "f (of_int n / of_int d) = 0" if "n < \r0 * of_int d\" for n proof - have "rat_of_int n < r0 * rat_of_int d" using that by linarith thus ?thesis using \d > 0\ by (intro r0) (auto simp: field_simps) qed have "eventually (\n. n > -\r0 * of_int d\) at_top" by (rule eventually_gt_at_top) hence "eventually (\n. f (of_int (-n) / of_int d) = 0) at_top" by (eventually_elim) (rule *, auto) hence "eventually (\n. f (of_int (-int n) / of_int d) = 0) at_top" by (rule eventually_compose_filterlim) (rule filterlim_int_sequentially) thus "eventually (\n. f (of_int (-int n) / of_int d) = 0) cofinite" by (simp add: cofinite_eq_sequentially) qed lemma fls_nth_of_fpxs: "fls_nth (fls_of_fpxs f) n = fpxs_nth f (of_int n / of_nat (fpxs_root_order f))" by transfer simp subsection \Basic algebraic typeclass instances\ instantiation fpxs :: (zero) zero begin lift_definition zero_fpxs :: "'a fpxs" is "\r::rat. 0 :: 'a" by (auto simp: is_fpxs_def supp_def) instance .. end instantiation fpxs :: ("{one, zero}") one begin lift_definition one_fpxs :: "'a fpxs" is "\r::rat. if r = 0 then 1 else 0 :: 'a" by (cases "(1 :: 'a) = 0") (auto simp: is_fpxs_def cong: if_cong) instance .. end lemma fls_of_fpxs_0 [simp]: "fls_of_fpxs 0 = 0" by transfer auto lemma fpxs_nth_0 [simp]: "fpxs_nth 0 r = 0" by transfer auto lemma fpxs_nth_1: "fpxs_nth 1 r = (if r = 0 then 1 else 0)" by transfer auto lemma fpxs_nth_1': "fpxs_nth 1 0 = 1" "r \ 0 \ fpxs_nth 1 r = 0" by (auto simp: fpxs_nth_1) instantiation fpxs :: (monoid_add) monoid_add begin lift_definition plus_fpxs :: "'a fpxs \ 'a fpxs \ 'a fpxs" is "\f g x. f x + g x" proof - fix f g :: "rat \ 'a" assume fg: "is_fpxs f" "is_fpxs g" show "is_fpxs (\x. f x + g x)" unfolding is_fpxs_def proof have supp: "supp (\x. f x + g x) \ supp f \ supp g" by (auto simp: supp_def) show "bdd_below (supp (\x. f x + g x))" by (rule bdd_below_mono[OF _ supp]) (use fg in \auto simp: is_fpxs_def\) have "(LCM r\supp (\x. f x + g x). snd (quotient_of r)) dvd (LCM r\supp f \ supp g. snd (quotient_of r))" by (intro Lcm_subset image_mono supp) also have "\ = lcm (LCM r\supp f. snd (quotient_of r)) (LCM r\supp g. snd (quotient_of r))" unfolding image_Un Lcm_Un .. finally have "(LCM r\supp (\x. f x + g x). snd (quotient_of r)) dvd lcm (LCM r\supp f. snd (quotient_of r)) (LCM r\supp g. snd (quotient_of r))" . moreover have "lcm (LCM r\supp f. snd (quotient_of r)) (LCM r\supp g. snd (quotient_of r)) \ 0" using fg by (auto simp: is_fpxs_def) ultimately show "(LCM r\supp (\x. f x + g x). snd (quotient_of r)) \ 0" by auto qed qed instance by standard (transfer; simp add: algebra_simps fun_eq_iff)+ end instance fpxs :: (comm_monoid_add) comm_monoid_add proof fix f g :: "'a fpxs" show "f + g = g + f" by transfer (auto simp: add_ac) qed simp_all lemma fpxs_nth_add [simp]: "fpxs_nth (f + g) r = fpxs_nth f r + fpxs_nth g r" by transfer auto lift_definition fpxs_of_fls :: "'a :: zero fls \ 'a fpxs" is "\f r. if r \ \ then f \r\ else 0" proof - fix f :: "int \ 'a" assume "eventually (\n. f (-int n) = 0) cofinite" hence "eventually (\n. f (-int n) = 0) at_top" by (simp add: cofinite_eq_sequentially) then obtain N where N: "f (-int n) = 0" if "n \ N" for n by (auto simp: eventually_at_top_linorder) show "is_fpxs (\r. if r \ \ then f \r\ else 0)" unfolding is_fpxs_def proof have "bdd_below {-(of_nat N::rat)..}" by simp moreover have "supp (\r::rat. if r \ \ then f \r\ else 0) \ {-of_nat N..}" proof fix r :: rat assume "r \ supp (\r. if r \ \ then f \r\ else 0)" then obtain m where [simp]: "r = of_int m" "f m \ 0" by (auto simp: supp_def elim!: Ints_cases split: if_splits) have "m \ -int N" using N[of "nat (-m)"] by (cases "m \ 0"; cases "-int N \ m") (auto simp: le_nat_iff) thus "r \ {-of_nat N..}" by simp qed ultimately show "bdd_below (supp (\r::rat. if r \ \ then f \r\ else 0))" by (rule bdd_below_mono) next have "(LCM r\supp (\r. if r \ \ then f \r\ else 0). snd (quotient_of r)) dvd 1" by (intro Lcm_least) (auto simp: supp_def elim!: Ints_cases split: if_splits) thus "(LCM r\supp (\r. if r \ \ then f \r\ else 0). snd (quotient_of r)) \ 0" by (intro notI) simp qed qed instantiation fpxs :: (group_add) group_add begin lift_definition uminus_fpxs :: "'a fpxs \ 'a fpxs" is "\f x. -f x" by (auto simp: is_fpxs_def) definition minus_fpxs :: "'a fpxs \ 'a fpxs \ 'a fpxs" where "minus_fpxs f g = f + (-g)" instance proof fix f :: "'a fpxs" show "-f + f = 0" by transfer auto qed (auto simp: minus_fpxs_def) end lemma fpxs_nth_uminus [simp]: "fpxs_nth (-f) r = -fpxs_nth f r" by transfer auto lemma fpxs_nth_minus [simp]: "fpxs_nth (f - g) r = fpxs_nth f r - fpxs_nth g r" unfolding minus_fpxs_def fpxs_nth_add fpxs_nth_uminus by simp lemma fpxs_of_fls_eq_iff [simp]: "fpxs_of_fls f = fpxs_of_fls g \ f = g" by transfer (force simp: fun_eq_iff Ints_def) lemma fpxs_of_fls_0 [simp]: "fpxs_of_fls 0 = 0" by transfer auto lemma fpxs_of_fls_1 [simp]: "fpxs_of_fls 1 = 1" by transfer (auto simp: fun_eq_iff elim!: Ints_cases) lemma fpxs_of_fls_add [simp]: "fpxs_of_fls (f + g) = fpxs_of_fls f + fpxs_of_fls g" by transfer (auto simp: fun_eq_iff elim!: Ints_cases) lemma fps_to_fls_sum [simp]: "fps_to_fls (sum f A) = (\x\A. fps_to_fls (f x))" by (induction A rule: infinite_finite_induct) auto lemma fpxs_of_fls_sum [simp]: "fpxs_of_fls (sum f A) = (\x\A. fpxs_of_fls (f x))" by (induction A rule: infinite_finite_induct) auto lemma fpxs_nth_of_fls: "fpxs_nth (fpxs_of_fls f) r = (if r \ \ then fls_nth f \r\ else 0)" by transfer auto lemma fpxs_of_fls_eq_0_iff [simp]: "fpxs_of_fls f = 0 \ f = 0" using fpxs_of_fls_eq_iff[of f 0] by (simp del: fpxs_of_fls_eq_iff) lemma fpxs_of_fls_eq_1_iff [simp]: "fpxs_of_fls f = 1 \ f = 1" using fpxs_of_fls_eq_iff[of f 1] by (simp del: fpxs_of_fls_eq_iff) lemma fpxs_root_order_of_fls [simp]: "fpxs_root_order (fpxs_of_fls f) = 1" proof (transfer, goal_cases) case (1 f) have "supp (\r. if r \ \ then f \r\ else 0) = rat_of_int ` {n. f n \ 0}" by (force simp: supp_def Ints_def) also have "(LCM r\\. snd (quotient_of r)) = nat (LCM x\{n. f n \ 0}. 1)" by (simp add: image_image) also have "\ = 1" by simp also have "nat 1 = 1" by simp finally show ?case . qed subsection \The substitution $X \mapsto X^r$\ text \ This operation turns a formal Puiseux series $f(X)$ into $f(X^r)$, where $r$ can be any positive rational number: \ lift_definition fpxs_compose_power :: "'a :: zero fpxs \ rat \ 'a fpxs" is "\f r x. if r > 0 then f (x / r) else 0" proof - fix f :: "rat \ 'a" and r :: rat assume f: "is_fpxs f" have "is_fpxs (\x. f (x / r))" if "r > 0" unfolding is_fpxs_def proof define r' where "r' = inverse r" have "r' > 0" using \r > 0\ by (auto simp: r'_def) have "(\x. x / r') ` supp f = supp (\x. f (x * r'))" using \r' > 0\ by (auto simp: supp_def image_iff vimage_def field_simps) hence eq: "(\x. x * r) ` supp f = supp (\x. f (x / r))" using \r > 0\ by (simp add: r'_def field_simps) from f have "bdd_below (supp f)" by (auto simp: is_fpxs_def) hence "bdd_below ((\x. x * r) ` supp f)" using \r > 0\ by (intro bdd_below_image_mono) (auto simp: mono_def divide_right_mono) also note eq finally show "bdd_below (supp (\x. f (x / r)))" . define a b where "a = fst (quotient_of r)" and "b = snd (quotient_of r)" have "b > 0" by (simp add: b_def quotient_of_denom_pos') have [simp]: "quotient_of r = (a, b)" by (simp add: a_def b_def) have "r = of_int a / of_int b" by (simp add: quotient_of_div) with \r > 0\ and \b > 0\ have \a > 0\ by (simp add: field_simps) have "(LCM r\supp (\x. f (x / r)). snd (quotient_of r)) = (LCM x\supp f. snd (quotient_of (x * r)))" by (simp add: eq [symmetric] image_image) also have "\ dvd (LCM x\supp f. snd (quotient_of x) * b)" using \a > 0\ \b > 0\ by (intro Lcm_mono) (simp add: rat_times_code case_prod_unfold Let_def Rat.normalize_def quotient_of_denom_pos' div_dvd_self) also have "\ dvd normalize (b * (LCM x\supp f. snd (quotient_of x)))" proof (cases "supp f = {}") case False thus ?thesis using Lcm_mult[of "(\x. snd (quotient_of x)) ` supp f" b] by (simp add: mult_ac image_image) qed auto hence "(LCM x\supp f. snd (quotient_of x) * b) dvd b * (LCM x\supp f. snd (quotient_of x))" by simp finally show "(LCM r\supp (\x. f (x / r)). snd (quotient_of r)) \ 0" using \b > 0\ f by (auto simp: is_fpxs_def) qed thus "is_fpxs (\x. if r > 0 then f (x / r) else 0)" by (cases "r > 0") (auto simp: is_fpxs_def supp_def) qed lemma fpxs_as_fls: "fpxs_compose_power (fpxs_of_fls (fls_of_fpxs f)) (1 / of_nat (fpxs_root_order f)) = f" proof (transfer, goal_cases) case (1 f) define d where "d = (LCM r\supp f. snd (quotient_of r))" have "d \ 0" by (simp add: d_def) moreover have "d \ 0" using 1 by (simp add: is_fpxs_def d_def) ultimately have "d > 0" by linarith have "(if rat_of_int d * x \ \ then f (rat_of_int \rat_of_int d * x\ / rat_of_int d) else 0) = f x" for x proof (cases "rat_of_int d * x \ \") case True then obtain n where n: "rat_of_int d * x = of_int n" by (auto elim!: Ints_cases) have "f (rat_of_int \rat_of_int d * x\ / rat_of_int d) = f (rat_of_int n / rat_of_int d)" by (simp add: n) also have "rat_of_int n / rat_of_int d = x" using n \d > 0\ by (simp add: field_simps) finally show ?thesis using True by simp next case False have "x \ supp f" proof assume "x \ supp f" hence "snd (quotient_of x) dvd d" by (simp add: d_def) hence "rat_of_int (fst (quotient_of x) * d) / rat_of_int (snd (quotient_of x)) \ \" by (intro of_int_divide_in_Ints) auto also have "rat_of_int (fst (quotient_of x) * d) / rat_of_int (snd (quotient_of x)) = rat_of_int d * (rat_of_int (fst (quotient_of x)) / rat_of_int (snd (quotient_of x)))" by (simp only: of_int_mult mult_ac times_divide_eq_right) also have "\ = rat_of_int d * x" by (metis Fract_of_int_quotient Rat_cases normalize_stable prod.sel(1) prod.sel(2) quotient_of_Fract) finally have "rat_of_int d * x \ \" . with False show False by contradiction qed thus ?thesis using False by (simp add: supp_def) qed thus ?case using \d > 0\ by (simp add: is_fpxs_def d_def mult_ac fun_eq_iff cong: if_cong) qed lemma fpxs_compose_power_0 [simp]: "fpxs_compose_power 0 r = 0" by transfer simp lemma fpxs_compose_power_1 [simp]: "r > 0 \ fpxs_compose_power 1 r = 1" by transfer (auto simp: fun_eq_iff) lemma fls_of_fpxs_eq_0_iff [simp]: "fls_of_fpxs x = 0 \ x = 0" by (metis fls_of_fpxs_0 fpxs_as_fls fpxs_compose_power_0 fpxs_of_fls_0) lemma fpxs_of_fls_compose_power [simp]: "fpxs_of_fls (fls_compose_power f d) = fpxs_compose_power (fpxs_of_fls f) (of_nat d)" proof (transfer, goal_cases) case (1 f d) show ?case proof (cases "d = 0") case False show ?thesis proof (intro ext, goal_cases) case (1 r) show ?case proof (cases "r \ \") case True then obtain n where [simp]: "r = of_int n" by (cases r rule: Ints_cases) show ?thesis proof (cases "d dvd n") case True thus ?thesis by (auto elim!: Ints_cases) next case False hence "rat_of_int n / rat_of_int (int d) \ \" using \d \ 0\ by (subst of_int_div_of_int_in_Ints_iff) auto thus ?thesis using False by auto qed next case False hence "r / rat_of_nat d \ \" using \d \ 0\ by (auto elim!: Ints_cases simp: field_simps) thus ?thesis using False by auto qed qed qed auto qed lemma fpxs_compose_power_add [simp]: "fpxs_compose_power (f + g) r = fpxs_compose_power f r + fpxs_compose_power g r" by transfer (auto simp: fun_eq_iff) lemma fpxs_compose_power_distrib: "r1 > 0 \ r2 > 0 \ fpxs_compose_power (fpxs_compose_power f r1) r2 = fpxs_compose_power f (r1 * r2)" by transfer (auto simp: fun_eq_iff algebra_simps zero_less_mult_iff) lemma fpxs_compose_power_divide_right: "r1 > 0 \ r2 > 0 \ fpxs_compose_power f (r1 / r2) = fpxs_compose_power (fpxs_compose_power f r1) (inverse r2)" by (simp add: fpxs_compose_power_distrib field_simps) lemma fpxs_compose_power_1_right [simp]: "fpxs_compose_power f 1 = f" by transfer auto lemma fpxs_compose_power_eq_iff [simp]: assumes "r > 0" shows "fpxs_compose_power f r = fpxs_compose_power g r \ f = g" using assms proof (transfer, goal_cases) case (1 r f g) have "f x = g x" if "\x. f (x / r) = g (x / r)" for x using that[of "x * r"] \r > 0\ by auto thus ?case using \r > 0\ by (auto simp: fun_eq_iff) qed lemma fpxs_compose_power_eq_1_iff [simp]: assumes "l > 0" shows "fpxs_compose_power p l = 1 \ p = 1" proof - have "fpxs_compose_power p l = 1 \ fpxs_compose_power p l = fpxs_compose_power 1 l" by (subst fpxs_compose_power_1) (use assms in auto) also have "\ \ p = 1" using assms by (subst fpxs_compose_power_eq_iff) auto finally show ?thesis . qed lemma fpxs_compose_power_eq_0_iff [simp]: assumes "r > 0" shows "fpxs_compose_power f r = 0 \ f = 0" using fpxs_compose_power_eq_iff[of r f 0] assms by (simp del: fpxs_compose_power_eq_iff) lemma fls_of_fpxs_of_fls [simp]: "fls_of_fpxs (fpxs_of_fls f) = f" using fpxs_as_fls[of "fpxs_of_fls f"] by simp lemma fpxs_as_fls': assumes "fpxs_root_order f dvd d" "d > 0" obtains f' where "f = fpxs_compose_power (fpxs_of_fls f') (1 / of_nat d)" proof - define D where "D = fpxs_root_order f" have "D > 0" by (auto simp: D_def) define f' where "f' = fls_of_fpxs f" from assms obtain d' where d': "d = D * d'" by (auto simp: D_def) have "d' > 0" using assms by (auto intro!: Nat.gr0I simp: d') define f'' where "f'' = fls_compose_power f' d'" have "fpxs_compose_power (fpxs_of_fls f'') (1 / of_nat d) = f" using \D > 0\ \d' > 0\ by (simp add: d' D_def f''_def f'_def fpxs_as_fls fpxs_compose_power_distrib) thus ?thesis using that[of f''] by blast qed subsection \Mutiplication and ring properties\ instantiation fpxs :: (comm_semiring_1) comm_semiring_1 begin lift_definition times_fpxs :: "'a fpxs \ 'a fpxs \ 'a fpxs" is "\f g x. (\(y,z) | y \ supp f \ z \ supp g \ x = y + z. f y * g z)" proof - fix f g :: "rat \ 'a" assume fg: "is_fpxs f" "is_fpxs g" show "is_fpxs (\x. \(y,z) | y \ supp f \ z \ supp g \ x = y + z. f y * g z)" (is "is_fpxs ?h") unfolding is_fpxs_def proof from fg obtain bnd1 bnd2 where bnds: "\x\supp f. x \ bnd1" "\x\supp g. x \ bnd2" by (auto simp: is_fpxs_def bdd_below_def) have "supp ?h \ (\(x,y). x + y) ` (supp f \ supp g)" proof fix x :: rat assume "x \ supp ?h" have "{(y,z). y \ supp f \ z \ supp g \ x = y + z} \ {}" proof assume eq: "{(y,z). y \ supp f \ z \ supp g \ x = y + z} = {}" hence "?h x = 0" by (simp only:) auto with \x \ supp ?h\ show False by (auto simp: supp_def) qed thus "x \ (\(x,y). x + y) ` (supp f \ supp g)" by auto qed also have "\ \ {bnd1 + bnd2..}" using bnds by (auto intro: add_mono) finally show "bdd_below (supp ?h)" by auto next define d1 where "d1 = (LCM r\supp f. snd (quotient_of r))" define d2 where "d2 = (LCM r\supp g. snd (quotient_of r))" have "(LCM r\supp ?h. snd (quotient_of r)) dvd (d1 * d2)" proof (intro Lcm_least, safe) fix r :: rat assume "r \ supp ?h" hence "(\(y, z) | y \ supp f \ z \ supp g \ r = y + z. f y * g z) \ 0" by (auto simp: supp_def) hence "{(y, z). y \ supp f \ z \ supp g \ r = y + z} \ {}" by (intro notI) simp_all then obtain y z where yz: "y \ supp f" "z \ supp g" "r = y + z" by auto have "snd (quotient_of r) = snd (quotient_of y) * snd (quotient_of z) div gcd (fst (quotient_of y) * snd (quotient_of z) + fst (quotient_of z) * snd (quotient_of y)) (snd (quotient_of y) * snd (quotient_of z))" by (simp add: \r = _\ rat_plus_code case_prod_unfold Let_def Rat.normalize_def quotient_of_denom_pos') also have "\ dvd snd (quotient_of y) * snd (quotient_of z)" by (metis dvd_def dvd_div_mult_self gcd_dvd2) also have "\ dvd d1 * d2" using yz by (auto simp: d1_def d2_def intro!: mult_dvd_mono) finally show "snd (quotient_of r) dvd d1 * d2" by (simp add: d1_def d2_def) qed moreover have "d1 * d2 \ 0" using fg by (auto simp: d1_def d2_def is_fpxs_def) ultimately show "(LCM r\supp ?h. snd (quotient_of r)) \ 0" by auto qed qed lemma fpxs_nth_mult: "fpxs_nth (f * g) r = (\(y,z) | y \ fpxs_supp f \ z \ fpxs_supp g \ r = y + z. fpxs_nth f y * fpxs_nth g z)" by transfer simp lemma fpxs_compose_power_mult [simp]: "fpxs_compose_power (f * g) r = fpxs_compose_power f r * fpxs_compose_power g r" proof (transfer, rule ext, goal_cases) case (1 f g r x) show ?case proof (cases "r > 0") case True have "(\x\{(y, z). y \ supp f \ z \ supp g \ x / r = y + z}. case x of (y, z) \ f y * g z) = (\x\{(y, z). y \ supp (\x. f (x / r)) \ z \ supp (\x. g (x / r)) \ x = y + z}. case x of (y, z) \ f (y / r) * g (z / r))" by (rule sum.reindex_bij_witness[of _ "\(x,y). (x/r,y/r)" "\(x,y). (x*r,y*r)"]) (use \r > 0\ in \auto simp: supp_def field_simps\) thus ?thesis by (auto simp: fun_eq_iff) qed auto qed lemma fpxs_supp_of_fls: "fpxs_supp (fpxs_of_fls f) = of_int ` supp (fls_nth f)" by (force simp: fpxs_supp_def fpxs_nth_of_fls supp_def elim!: Ints_cases) lemma fpxs_of_fls_mult [simp]: "fpxs_of_fls (f * g) = fpxs_of_fls f * fpxs_of_fls g" proof (rule fpxs_ext) fix r :: rat show "fpxs_nth (fpxs_of_fls (f * g)) r = fpxs_nth (fpxs_of_fls f * fpxs_of_fls g) r" proof (cases "r \ \") case True define h1 where "h1 = (\(x, y). (\x::rat\, \y::rat\))" define h2 where "h2 = (\(x, y). (of_int x :: rat, of_int y :: rat))" define df dg where [simp]: "df = fls_subdegree f" "dg = fls_subdegree g" from True obtain n where [simp]: "r = of_int n" by (cases rule: Ints_cases) have "fpxs_nth (fpxs_of_fls f * fpxs_of_fls g) r = (\(y,z) | y \ fpxs_supp (fpxs_of_fls f) \ z \ fpxs_supp (fpxs_of_fls g) \ rat_of_int n = y + z. (if y \ \ then fls_nth f \y\ else 0) * (if z \ \ then fls_nth g \z\ else 0))" by (auto simp: fpxs_nth_mult fpxs_nth_of_fls) also have "\ = (\(y,z) | y \ supp (fls_nth f) \ z \ supp (fls_nth g) \ n = y + z. fls_nth f y * fls_nth g z)" by (rule sum.reindex_bij_witness[of _ h2 h1]) (auto simp: h1_def h2_def fpxs_supp_of_fls) also have "\ = (\y | y - fls_subdegree g \ supp (fls_nth f) \ fls_subdegree g + n - y \ supp (fls_nth g). fls_nth f (y - fls_subdegree g) * fls_nth g (fls_subdegree g + n - y))" by (rule sum.reindex_bij_witness[of _ "\y. (y - fls_subdegree g, fls_subdegree g + n - y)" "\z. fst z + fls_subdegree g"]) auto also have "\ = (\i = fls_subdegree f + fls_subdegree g..n. fls_nth f (i - fls_subdegree g) * fls_nth g (fls_subdegree g + n - i))" using fls_subdegree_leI[of f] fls_subdegree_leI [of g] by (intro sum.mono_neutral_left; force simp: supp_def) also have "\ = fpxs_nth (fpxs_of_fls (f * g)) r" by (auto simp: fls_times_nth fpxs_nth_of_fls) finally show ?thesis .. next case False have "fpxs_nth (fpxs_of_fls f * fpxs_of_fls g) r = (\(y,z) | y \ fpxs_supp (fpxs_of_fls f) \ z \ fpxs_supp (fpxs_of_fls g) \ r = y + z. (if y \ \ then fls_nth f \y\ else 0) * (if z \ \ then fls_nth g \z\ else 0))" by (simp add: fpxs_nth_mult fpxs_nth_of_fls) also have "\ = 0" using False by (intro sum.neutral ballI) auto also have "0 = fpxs_nth (fpxs_of_fls (f * g)) r" using False by (simp add: fpxs_nth_of_fls) finally show ?thesis .. qed qed instance proof show "0 \ (1 :: 'a fpxs)" by transfer (auto simp: fun_eq_iff) next fix f :: "'a fpxs" show "1 * f = f" proof (transfer, goal_cases) case (1 f) have "{(y, z). y \ supp (\r. if r = 0 then (1::'a) else 0) \ z \ supp f \ x = y + z} = (if x \ supp f then {(0, x)} else {})" for x by (auto simp: supp_def split: if_splits) thus ?case by (auto simp: fun_eq_iff supp_def) qed next fix f :: "'a fpxs" show "0 * f = 0" by transfer (auto simp: fun_eq_iff supp_def) show "f * 0 = 0" by transfer (auto simp: fun_eq_iff supp_def) next fix f g :: "'a fpxs" show "f * g = g * f" proof (transfer, rule ext, goal_cases) case (1 f g x) show "(\(y, z)\{(y, z). y \ supp f \ z \ supp g \ x = y + z}. f y * g z) = (\(y, z)\{(y, z). y \ supp g \ z \ supp f \ x = y + z}. g y * f z)" by (rule sum.reindex_bij_witness[of _ "\(x,y). (y,x)" "\(x,y). (y,x)"]) (auto simp: mult_ac) qed next fix f g h :: "'a fpxs" define d where "d = (LCM F\{f,g,h}. fpxs_root_order F)" have "d > 0" by (auto simp: d_def intro!: Nat.gr0I) obtain f' where f: "f = fpxs_compose_power (fpxs_of_fls f') (1 / of_nat d)" using fpxs_as_fls'[of f d] \d > 0\ by (auto simp: d_def) obtain g' where g: "g = fpxs_compose_power (fpxs_of_fls g') (1 / of_nat d)" using fpxs_as_fls'[of g d] \d > 0\ by (auto simp: d_def) obtain h' where h: "h = fpxs_compose_power (fpxs_of_fls h') (1 / of_nat d)" using fpxs_as_fls'[of h d] \d > 0\ by (auto simp: d_def) show "(f * g) * h = f * (g * h)" by (simp add: f g h mult_ac flip: fpxs_compose_power_mult fpxs_compose_power_add fpxs_of_fls_mult) show "(f + g) * h = f * h + g * h" by (simp add: f g h ring_distribs flip: fpxs_compose_power_mult fpxs_compose_power_add fpxs_of_fls_mult fpxs_of_fls_add) qed end instance fpxs :: (comm_ring_1) comm_ring_1 by intro_classes auto instance fpxs :: ("{comm_semiring_1,semiring_no_zero_divisors}") semiring_no_zero_divisors proof fix f g :: "'a fpxs" assume fg: "f \ 0" "g \ 0" define d where "d = lcm (fpxs_root_order f) (fpxs_root_order g)" have "d > 0" by (auto simp: d_def intro!: lcm_pos_nat) obtain f' where f: "f = fpxs_compose_power (fpxs_of_fls f') (1 / of_nat d)" using fpxs_as_fls'[of f d] \d > 0\ by (auto simp: d_def) obtain g' where g: "g = fpxs_compose_power (fpxs_of_fls g') (1 / of_nat d)" using fpxs_as_fls'[of g d] \d > 0\ by (auto simp: d_def) show "f * g \ 0" using \d > 0\ fg by (simp add: f g flip: fpxs_compose_power_mult fpxs_of_fls_mult) qed lemma fpxs_of_fls_power [simp]: "fpxs_of_fls (f ^ n) = fpxs_of_fls f ^ n" by (induction n) auto lemma fpxs_compose_power_power [simp]: "r > 0 \ fpxs_compose_power (f ^ n) r = fpxs_compose_power f r ^ n" by (induction n) simp_all subsection \Constant Puiseux series and the series \X\\ lift_definition fpxs_const :: "'a :: zero \ 'a fpxs" is "\c n. if n = 0 then c else 0" proof - fix c :: 'a have "supp (\n::rat. if n = 0 then c else 0) = (if c = 0 then {} else {0})" by auto thus "is_fpxs (\n::rat. if n = 0 then c else 0)" unfolding is_fpxs_def by auto qed lemma fpxs_const_0 [simp]: "fpxs_const 0 = 0" by transfer auto lemma fpxs_const_1 [simp]: "fpxs_const 1 = 1" by transfer auto lemma fpxs_of_fls_const [simp]: "fpxs_of_fls (fls_const c) = fpxs_const c" by transfer (auto simp: fun_eq_iff Ints_def) lemma fls_of_fpxs_const [simp]: "fls_of_fpxs (fpxs_const c) = fls_const c" by (metis fls_of_fpxs_of_fls fpxs_of_fls_const) lemma fls_of_fpxs_1 [simp]: "fls_of_fpxs 1 = 1" using fls_of_fpxs_const[of 1] by (simp del: fls_of_fpxs_const) lift_definition fpxs_X :: "'a :: {one, zero} fpxs" is "\x. if x = 1 then (1::'a) else 0" by (cases "1 = (0 :: 'a)") (auto simp: is_fpxs_def cong: if_cong) lemma fpxs_const_altdef: "fpxs_const x = fpxs_of_fls (fls_const x)" by transfer auto lemma fpxs_const_add [simp]: "fpxs_const (x + y) = fpxs_const x + fpxs_const y" by transfer auto lemma fpxs_const_mult [simp]: fixes x y :: "'a::{comm_semiring_1}" shows "fpxs_const (x * y) = fpxs_const x * fpxs_const y" unfolding fpxs_const_altdef fls_const_mult_const[symmetric] fpxs_of_fls_mult .. lemma fpxs_const_eq_iff [simp]: "fpxs_const x = fpxs_const y \ x = y" by transfer (auto simp: fun_eq_iff) lemma of_nat_fpxs_eq: "of_nat n = fpxs_const (of_nat n)" by (induction n) auto lemma fpxs_const_uminus [simp]: "fpxs_const (-x) = -fpxs_const x" by transfer auto lemma fpxs_const_diff [simp]: "fpxs_const (x - y) = fpxs_const x - fpxs_const y" unfolding minus_fpxs_def by transfer auto lemma of_int_fpxs_eq: "of_int n = fpxs_const (of_int n)" by (induction n) (auto simp: of_nat_fpxs_eq) subsection \More algebraic typeclass instances\ instance fpxs :: ("{comm_semiring_1,semiring_char_0}") semiring_char_0 proof show "inj (of_nat :: nat \ 'a fpxs)" by (intro injI) (auto simp: of_nat_fpxs_eq) qed instance fpxs :: ("{comm_ring_1,ring_char_0}") ring_char_0 .. instance fpxs :: (idom) idom .. instantiation fpxs :: (field) field begin definition inverse_fpxs :: "'a fpxs \ 'a fpxs" where "inverse_fpxs f = fpxs_compose_power (fpxs_of_fls (inverse (fls_of_fpxs f))) (1 / of_nat (fpxs_root_order f))" definition divide_fpxs :: "'a fpxs \ 'a fpxs \ 'a fpxs" where "divide_fpxs f g = f * inverse g" instance proof fix f :: "'a fpxs" assume "f \ 0" define f' where "f' = fls_of_fpxs f" define d where "d = fpxs_root_order f" have "d > 0" by (auto simp: d_def) have f: "f = fpxs_compose_power (fpxs_of_fls f') (1 / of_nat d)" by (simp add: f'_def d_def fpxs_as_fls) have "inverse f * f = fpxs_compose_power (fpxs_of_fls (inverse f')) (1 / of_nat d) * f" by (simp add: inverse_fpxs_def f'_def d_def) also have "fpxs_compose_power (fpxs_of_fls (inverse f')) (1 / of_nat d) * f = fpxs_compose_power (fpxs_of_fls (inverse f' * f')) (1 / of_nat d)" by (simp add: f) also have "inverse f' * f' = 1" using \f \ 0\ \d > 0\ by (simp add: f field_simps) finally show "inverse f * f = 1" using \d > 0\ by simp qed (auto simp: divide_fpxs_def inverse_fpxs_def) end instance fpxs :: (field_char_0) field_char_0 .. subsection \Valuation\ definition fpxs_val :: "'a :: zero fpxs \ rat" where "fpxs_val f = of_int (fls_subdegree (fls_of_fpxs f)) / rat_of_nat (fpxs_root_order f)" lemma fpxs_val_of_fls [simp]: "fpxs_val (fpxs_of_fls f) = of_int (fls_subdegree f)" by (simp add: fpxs_val_def) lemma fpxs_nth_compose_power [simp]: assumes "r > 0" shows "fpxs_nth (fpxs_compose_power f r) n = fpxs_nth f (n / r)" using assms by transfer auto lemma fls_of_fpxs_uminus [simp]: "fls_of_fpxs (-f) = -fls_of_fpxs f" by transfer auto lemma fpxs_root_order_uminus [simp]: "fpxs_root_order (-f) = fpxs_root_order f" by transfer auto lemma fpxs_val_uminus [simp]: "fpxs_val (-f) = fpxs_val f" unfolding fpxs_val_def by simp lemma fpxs_val_minus_commute: "fpxs_val (f - g) = fpxs_val (g - f)" by (subst fpxs_val_uminus [symmetric]) (simp del: fpxs_val_uminus) lemma fpxs_val_const [simp]: "fpxs_val (fpxs_const c) = 0" by (simp add: fpxs_val_def) lemma fpxs_val_1 [simp]: "fpxs_val 1 = 0" by (simp add: fpxs_val_def) lemma of_int_fls_subdegree_of_fpxs: "rat_of_int (fls_subdegree (fls_of_fpxs f)) = fpxs_val f * of_nat (fpxs_root_order f)" by (simp add: fpxs_val_def) lemma fpxs_nth_val_nonzero: assumes "f \ 0" shows "fpxs_nth f (fpxs_val f) \ 0" proof - define N where "N = fpxs_root_order f" define f' where "f' = fls_of_fpxs f" define M where "M = fls_subdegree f'" have val: "fpxs_val f = of_int M / of_nat N" by (simp add: M_def fpxs_val_def N_def f'_def) have *: "f = fpxs_compose_power (fpxs_of_fls f') (1 / rat_of_nat N)" by (simp add: fpxs_as_fls N_def f'_def) also have "fpxs_nth \ (fpxs_val f) = fpxs_nth (fpxs_of_fls f') (fpxs_val f * rat_of_nat (fpxs_root_order f))" by (subst fpxs_nth_compose_power) (auto simp: N_def) also have "\ = fls_nth f' M" by (subst fpxs_nth_of_fls) (auto simp: val N_def) also have "f' \ 0" using * assms by auto hence "fls_nth f' M \ 0" unfolding M_def by simp finally show "fpxs_nth f (fpxs_val f) \ 0" . qed lemma fpxs_nth_below_val: assumes n: "n < fpxs_val f" shows "fpxs_nth f n = 0" proof (cases "f = 0") case False define N where "N = fpxs_root_order f" define f' where "f' = fls_of_fpxs f" define M where "M = fls_subdegree f'" have val: "fpxs_val f = of_int M / of_nat N" by (simp add: M_def fpxs_val_def N_def f'_def) have *: "f = fpxs_compose_power (fpxs_of_fls f') (1 / rat_of_nat N)" by (simp add: fpxs_as_fls N_def f'_def) have "fpxs_nth f n = fpxs_nth (fpxs_of_fls f') (n * rat_of_nat N)" by (subst *, subst fpxs_nth_compose_power) (auto simp: N_def) also have "\ = 0" proof (cases "rat_of_nat N * n \ \") case True then obtain n' where n': "of_int n' = rat_of_nat N * n" by (elim Ints_cases) auto have "of_int n' < rat_of_nat N * fpxs_val f" unfolding n' using n by (intro mult_strict_left_mono) (auto simp: N_def) also have "\ = of_int M" by (simp add: val N_def) finally have "n' < M" by linarith have "fpxs_nth (fpxs_of_fls f') (rat_of_nat N * n) = fls_nth f' n'" unfolding n'[symmetric] by (subst fpxs_nth_of_fls) (auto simp: N_def) also from \n' < M\ have "\ = 0" unfolding M_def by simp finally show ?thesis by (simp add: mult_ac) qed (auto simp: fpxs_nth_of_fls mult_ac) finally show "fpxs_nth f n = 0" . qed auto lemma fpxs_val_leI: "fpxs_nth f r \ 0 \ fpxs_val f \ r" using fpxs_nth_below_val[of r f] by (cases "f = 0"; cases "fpxs_val f" r rule: linorder_cases) auto lemma fpxs_val_0 [simp]: "fpxs_val 0 = 0" by (simp add: fpxs_val_def) lemma fpxs_val_geI: assumes "f \ 0" "\r. r < r' \ fpxs_nth f r = 0" shows "fpxs_val f \ r'" using fpxs_nth_val_nonzero[of f] assms by force lemma fpxs_val_compose_power [simp]: assumes "r > 0" shows "fpxs_val (fpxs_compose_power f r) = fpxs_val f * r" proof (cases "f = 0") case [simp]: False show ?thesis proof (intro antisym) show "fpxs_val (fpxs_compose_power f r) \ fpxs_val f * r" using assms by (intro fpxs_val_leI) (simp add: fpxs_nth_val_nonzero) next show "fpxs_val f * r \ fpxs_val (fpxs_compose_power f r)" proof (intro fpxs_val_geI) show "fpxs_nth (fpxs_compose_power f r) r' = 0" if "r' < fpxs_val f * r" for r' unfolding fpxs_nth_compose_power[OF assms] by (rule fpxs_nth_below_val) (use that assms in \auto simp: field_simps\) qed (use assms in auto) qed qed auto lemma fpxs_val_add_ge: assumes "f + g \ 0" shows "fpxs_val (f + g) \ min (fpxs_val f) (fpxs_val g)" proof (rule ccontr) assume "\(fpxs_val (f + g) \ min (fpxs_val f) (fpxs_val g))" (is "\(?n \ _)") hence "?n < fpxs_val f" "?n < fpxs_val g" by auto hence "fpxs_nth f ?n = 0" "fpxs_nth g ?n = 0" by (intro fpxs_nth_below_val; simp; fail)+ hence "fpxs_nth (f + g) ?n = 0" by simp moreover have "fpxs_nth (f + g) ?n \ 0" by (intro fpxs_nth_val_nonzero assms) ultimately show False by contradiction qed lemma fpxs_val_diff_ge: assumes "f \ g" shows "fpxs_val (f - g) \ min (fpxs_val f) (fpxs_val g)" using fpxs_val_add_ge[of f "-g"] assms by simp lemma fpxs_nth_mult_val: "fpxs_nth (f * g) (fpxs_val f + fpxs_val g) = fpxs_nth f (fpxs_val f) * fpxs_nth g (fpxs_val g)" proof (cases "f = 0 \ g = 0") case False have "{(y, z). y \ fpxs_supp f \ z \ fpxs_supp g \ fpxs_val f + fpxs_val g = y + z} \ {(fpxs_val f, fpxs_val g)}" using False fpxs_val_leI[of f] fpxs_val_leI[of g] by (force simp: fpxs_supp_def supp_def) hence "fpxs_nth (f * g) (fpxs_val f + fpxs_val g) = (\(y, z)\{(fpxs_val f, fpxs_val g)}. fpxs_nth f y * fpxs_nth g z)" unfolding fpxs_nth_mult by (intro sum.mono_neutral_left) (auto simp: fpxs_supp_def supp_def) thus ?thesis by simp qed auto lemma fpxs_val_mult [simp]: fixes f g :: "'a :: {comm_semiring_1, semiring_no_zero_divisors} fpxs" assumes "f \ 0" "g \ 0" shows "fpxs_val (f * g) = fpxs_val f + fpxs_val g" proof (intro antisym fpxs_val_leI fpxs_val_geI) fix r :: rat assume r: "r < fpxs_val f + fpxs_val g" show "fpxs_nth (f * g) r = 0" unfolding fpxs_nth_mult using assms fpxs_val_leI[of f] fpxs_val_leI[of g] r by (intro sum.neutral; force) qed (use assms in \auto simp: fpxs_nth_mult_val fpxs_nth_val_nonzero\) lemma fpxs_val_power [simp]: fixes f :: "'a :: {comm_semiring_1, semiring_no_zero_divisors} fpxs" assumes "f \ 0 \ n > 0" shows "fpxs_val (f ^ n) = of_nat n * fpxs_val f" proof (cases "f = 0") case False have [simp]: "f ^ n \ 0" for n using False by (induction n) auto thus ?thesis using False by (induction n) (auto simp: algebra_simps) qed (use assms in \auto simp: power_0_left\) lemma fpxs_nth_power_val [simp]: fixes f :: "'a :: {comm_semiring_1, semiring_no_zero_divisors} fpxs" shows "fpxs_nth (f ^ r) (rat_of_nat r * fpxs_val f) = fpxs_nth f (fpxs_val f) ^ r" proof (cases "f \ 0") case True show ?thesis proof (induction r) case (Suc r) have "fpxs_nth (f ^ Suc r) (rat_of_nat (Suc r) * fpxs_val f) = fpxs_nth (f * f ^ r) (fpxs_val f + fpxs_val (f ^ r))" using True by (simp add: fpxs_nth_mult_val ring_distribs) also have "\ = fpxs_nth f (fpxs_val f) ^ Suc r" using Suc True by (subst fpxs_nth_mult_val) auto finally show ?case . qed (auto simp: fpxs_nth_1') next case False thus ?thesis by (cases r) (auto simp: fpxs_nth_1') qed subsection \Powers of \X\ and shifting\ lift_definition fpxs_X_power :: "rat \ 'a :: {zero, one} fpxs" is "\r n :: rat. if n = r then 1 else (0 :: 'a)" proof - fix r :: rat have "supp (\n. if n = r then 1 else (0 :: 'a)) = (if (1 :: 'a) = 0 then {} else {r})" by (auto simp: supp_def) thus "is_fpxs (\n. if n = r then 1 else (0 :: 'a))" using quotient_of_denom_pos'[of r] by (auto simp: is_fpxs_def) qed lemma fpxs_X_power_0 [simp]: "fpxs_X_power 0 = 1" by transfer auto lemma fpxs_X_power_add: "fpxs_X_power (a + b) = fpxs_X_power a * fpxs_X_power b" proof (transfer, goal_cases) case (1 a b) have *: "{(y,z). y \ supp (\n. if n=a then (1::'a) else 0) \ z \ supp (\n. if n=b then (1::'a) else 0) \ x=y+z} = (if x = a + b then {(a, b)} else {})" for x by (auto simp: supp_def fun_eq_iff) show ?case unfolding * by (auto simp: fun_eq_iff case_prod_unfold) qed lemma fpxs_X_power_mult: "fpxs_X_power (rat_of_nat n * m) = fpxs_X_power m ^ n" by (induction n) (auto simp: ring_distribs fpxs_X_power_add) lemma fpxs_of_fls_X_power [simp]: "fpxs_of_fls (fls_shift n 1) = fpxs_X_power (-rat_of_int n)" by transfer (auto simp: fun_eq_iff Ints_def simp flip: of_int_minus) lemma fpxs_X_power_neq_0 [simp]: "fpxs_X_power r \ (0 :: 'a :: zero_neq_one fpxs)" by transfer (auto simp: fun_eq_iff) lemma fpxs_X_power_eq_1_iff [simp]: "fpxs_X_power r = (1 :: 'a :: zero_neq_one fpxs) \ r = 0" by transfer (auto simp: fun_eq_iff) lift_definition fpxs_shift :: "rat \ 'a :: zero fpxs \ 'a fpxs" is "\r f n. f (n + r)" proof - fix r :: rat and f :: "rat \ 'a" assume f: "is_fpxs f" have subset: "supp (\n. f (n + r)) \ (\n. n + r) -` supp f" by (auto simp: supp_def) have eq: "(\n. n + r) -` supp f = (\n. n - r) ` supp f" by (auto simp: image_iff algebra_simps) show "is_fpxs (\n. f (n + r))" unfolding is_fpxs_def proof have "bdd_below ((\n. n + r) -` supp f)" unfolding eq by (rule bdd_below_image_mono) (use f in \auto simp: is_fpxs_def mono_def\) thus "bdd_below (supp (\n. f (n + r)))" by (rule bdd_below_mono[OF _ subset]) next have "(LCM r\supp (\n. f (n + r)). snd (quotient_of r)) dvd (LCM r\(\n. n + r) -` supp f. snd (quotient_of r))" by (intro Lcm_subset image_mono subset) also have "\ = (LCM x\supp f. snd (quotient_of (x - r)))" by (simp only: eq image_image o_def) also have "\ dvd (LCM x\supp f. snd (quotient_of r) * snd (quotient_of x))" by (subst mult.commute, intro Lcm_mono quotient_of_denom_diff_dvd) also have "\ = Lcm ((\x. snd (quotient_of r) * x) ` (\x. snd (quotient_of x)) ` supp f)" by (simp add: image_image o_def) also have "\ dvd normalize (snd (quotient_of r) * (LCM x\supp f. snd (quotient_of x)))" proof (cases "supp f = {}") case False thus ?thesis by (subst Lcm_mult) auto qed auto finally show "(LCM r\supp (\n. f (n + r)). snd (quotient_of r)) \ 0" using quotient_of_denom_pos'[of r] f by (auto simp: is_fpxs_def) qed qed lemma fpxs_nth_shift [simp]: "fpxs_nth (fpxs_shift r f) n = fpxs_nth f (n + r)" by transfer simp_all lemma fpxs_shift_0_left [simp]: "fpxs_shift 0 f = f" by transfer auto lemma fpxs_shift_add_left: "fpxs_shift (m + n) f = fpxs_shift m (fpxs_shift n f)" by transfer (simp_all add: add_ac) lemma fpxs_shift_diff_left: "fpxs_shift (m - n) f = fpxs_shift m (fpxs_shift (-n) f)" by (subst fpxs_shift_add_left [symmetric]) auto lemma fpxs_shift_0 [simp]: "fpxs_shift r 0 = 0" by transfer simp_all lemma fpxs_shift_add [simp]: "fpxs_shift r (f + g) = fpxs_shift r f + fpxs_shift r g" by transfer auto lemma fpxs_shift_uminus [simp]: "fpxs_shift r (-f) = -fpxs_shift r f" by transfer auto lemma fpxs_shift_shift_uminus [simp]: "fpxs_shift r (fpxs_shift (-r) f) = f" by (simp flip: fpxs_shift_add_left) lemma fpxs_shift_shift_uminus' [simp]: "fpxs_shift (-r) (fpxs_shift r f) = f" by (simp flip: fpxs_shift_add_left) lemma fpxs_shift_diff [simp]: "fpxs_shift r (f - g) = fpxs_shift r f - fpxs_shift r g" unfolding minus_fpxs_def by (subst fpxs_shift_add) auto lemma fpxs_shift_compose_power [simp]: "fpxs_shift r (fpxs_compose_power f s) = fpxs_compose_power (fpxs_shift (r / s) f) s" by transfer (simp_all add: add_divide_distrib add_ac cong: if_cong) lemma rat_of_int_div_dvd: "d dvd n \ rat_of_int (n div d) = rat_of_int n / rat_of_int d" by auto lemma fpxs_of_fls_shift [simp]: "fpxs_of_fls (fls_shift n f) = fpxs_shift (of_int n) (fpxs_of_fls f)" proof (transfer, goal_cases) case (1 n f) show ?case proof fix r :: rat have eq: "r + rat_of_int n \ \ \ r \ \" by (metis Ints_add Ints_diff Ints_of_int add_diff_cancel_right') show "(if r \ \ then f (\r\ + n) else 0) = (if r + rat_of_int n \ \ then f \r + rat_of_int n\ else 0)" unfolding eq by auto qed qed lemma fpxs_shift_mult: "f * fpxs_shift r g = fpxs_shift r (f * g)" "fpxs_shift r f * g = fpxs_shift r (f * g)" proof - obtain a b where ab: "r = of_int a / of_nat b" and "b > 0" by (metis Fract_of_int_quotient of_int_of_nat_eq quotient_of_unique zero_less_imp_eq_int) define s where "s = lcm b (lcm (fpxs_root_order f) (fpxs_root_order g))" have "s > 0" using \b > 0\ by (auto simp: s_def intro!: Nat.gr0I) obtain f' where f: "f = fpxs_compose_power (fpxs_of_fls f') (1 / rat_of_nat s)" using fpxs_as_fls'[of f s] \s > 0\ by (auto simp: s_def) obtain g' where g: "g = fpxs_compose_power (fpxs_of_fls g') (1 / rat_of_nat s)" using fpxs_as_fls'[of g s] \s > 0\ by (auto simp: s_def) define n where "n = (a * s) div b" have "b dvd s" by (auto simp: s_def) have sr_eq: "r * rat_of_nat s = rat_of_int n" using \b > 0\ \b dvd s\ by (simp add: ab field_simps of_rat_divide of_rat_mult n_def rat_of_int_div_dvd) show "f * fpxs_shift r g = fpxs_shift r (f * g)" "fpxs_shift r f * g = fpxs_shift r (f * g)" unfolding f g using \s > 0\ by (simp_all flip: fpxs_compose_power_mult fpxs_of_fls_mult fpxs_of_fls_shift add: sr_eq fls_shifted_times_simps mult_ac) qed lemma fpxs_shift_1: "fpxs_shift r 1 = fpxs_X_power (-r)" by transfer (auto simp: fun_eq_iff) lemma fpxs_X_power_conv_shift: "fpxs_X_power r = fpxs_shift (-r) 1" by (simp add: fpxs_shift_1) lemma fpxs_shift_power [simp]: "fpxs_shift n x ^ m = fpxs_shift (of_nat m * n) (x ^ m)" by (induction m) (simp_all add: algebra_simps fpxs_shift_mult flip: fpxs_shift_add_left) lemma fpxs_compose_power_X_power [simp]: "s > 0 \ fpxs_compose_power (fpxs_X_power r) s = fpxs_X_power (r * s)" by transfer (simp add: field_simps) subsection \The \n\-th root of a Puiseux series\ text \ In this section, we define the formal root of a Puiseux series. This is done using the same concept for formal power series. There is still one interesting theorems that is missing here, e.g.\ the uniqueness (which could probably be lifted over from FPSs) somehow. \ definition fpxs_radical :: "(nat \ 'a :: field_char_0 \ 'a) \ nat \ 'a fpxs \ 'a fpxs" where "fpxs_radical rt r f = (if f = 0 then 0 else (let f' = fls_base_factor_to_fps (fls_of_fpxs f); f'' = fpxs_of_fls (fps_to_fls (fps_radical rt r f')) in fpxs_shift (-fpxs_val f / rat_of_nat r) (fpxs_compose_power f'' (1 / rat_of_nat (fpxs_root_order f)))))" lemma fpxs_radical_0 [simp]: "fpxs_radical rt r 0 = 0" by (simp add: fpxs_radical_def) lemma fixes r :: nat assumes r: "r > 0" shows fpxs_power_radical: "rt r (fpxs_nth f (fpxs_val f)) ^ r = fpxs_nth f (fpxs_val f) \ fpxs_radical rt r f ^ r = f" and fpxs_radical_lead_coeff: "f \ 0 \ fpxs_nth (fpxs_radical rt r f) (fpxs_val f / rat_of_nat r) = rt r (fpxs_nth f (fpxs_val f))" proof - define q where "q = fpxs_root_order f" define f' where "f' = fls_base_factor_to_fps (fls_of_fpxs f)" have [simp]: "fps_nth f' 0 = fpxs_nth f (fpxs_val f)" by (simp add: f'_def fls_nth_of_fpxs of_int_fls_subdegree_of_fpxs) define f'' where "f'' = fpxs_of_fls (fps_to_fls (fps_radical rt r f'))" have eq1: "fls_of_fpxs f = fls_shift (-fls_subdegree (fls_of_fpxs f)) (fps_to_fls f')" by (subst fls_conv_base_factor_to_fps_shift_subdegree) (simp add: f'_def) have eq2: "fpxs_compose_power (fpxs_of_fls (fls_of_fpxs f)) (1 / of_nat q) = f" unfolding q_def by (rule fpxs_as_fls) also note eq1 also have "fpxs_of_fls (fls_shift (- fls_subdegree (fls_of_fpxs f)) (fps_to_fls f')) = fpxs_shift (- (fpxs_val f * rat_of_nat q)) (fpxs_of_fls (fps_to_fls f'))" by (simp add: of_int_fls_subdegree_of_fpxs q_def) finally have eq3: "fpxs_compose_power (fpxs_shift (- (fpxs_val f * rat_of_nat q)) (fpxs_of_fls (fps_to_fls f'))) (1 / rat_of_nat q) = f" . { assume rt: "rt r (fpxs_nth f (fpxs_val f)) ^ r = fpxs_nth f (fpxs_val f)" show "fpxs_radical rt r f ^ r = f" proof (cases "f = 0") case [simp]: False have "f'' ^ r = fpxs_of_fls (fps_to_fls (fps_radical rt r f' ^ r))" by (simp add: fps_to_fls_power f''_def) also have "fps_radical rt r f' ^ r = f'" using power_radical[of f' rt "r - 1"] r rt by (simp add: fpxs_nth_val_nonzero) finally have "f'' ^ r = fpxs_of_fls (fps_to_fls f')" . have "fpxs_shift (-fpxs_val f / rat_of_nat r) (fpxs_compose_power f'' (1 / of_nat q)) ^ r = fpxs_shift (-fpxs_val f) (fpxs_compose_power (f'' ^ r) (1 / of_nat q))" unfolding q_def using r by (subst fpxs_shift_power, subst fpxs_compose_power_power [symmetric]) simp_all also have "f'' ^ r = fpxs_of_fls (fps_to_fls f')" by fact also have "fpxs_shift (-fpxs_val f) (fpxs_compose_power (fpxs_of_fls (fps_to_fls f')) (1 / of_nat q)) = f" using r eq3 by simp finally show "fpxs_radical rt r f ^ r = f" by (simp add: fpxs_radical_def f'_def f''_def q_def) qed (use r in auto) } assume [simp]: "f \ 0" have "fpxs_nth (fpxs_shift (-fpxs_val f / of_nat r) (fpxs_compose_power f'' (1 / of_nat q))) (fpxs_val f / of_nat r) = fpxs_nth f'' 0" using r by (simp add: q_def) also have "fpxs_shift (-fpxs_val f / of_nat r) (fpxs_compose_power f'' (1 / of_nat q)) = fpxs_radical rt r f" by (simp add: fpxs_radical_def q_def f'_def f''_def) also have "fpxs_nth f'' 0 = rt r (fpxs_nth f (fpxs_val f))" using r by (simp add: f''_def fpxs_nth_of_fls) finally show "fpxs_nth (fpxs_radical rt r f) (fpxs_val f / rat_of_nat r) = rt r (fpxs_nth f (fpxs_val f))" . qed lemma fls_base_factor_power: fixes f :: "'a::{semiring_1, semiring_no_zero_divisors} fls" shows "fls_base_factor (f ^ n) = fls_base_factor f ^ n" proof (cases "f = 0") case False have [simp]: "f ^ n \ 0" for n by (induction n) (use False in auto) show ?thesis using False by (induction n) (auto simp: fls_base_factor_mult simp flip: fls_times_both_shifted_simp) qed (cases n; simp) (* TODO: Uniqueness of radical. Also: composition and composition inverse *) hide_const (open) supp subsection \Algebraic closedness\ text \ We will now show that the field of formal Puiseux series over an algebraically closed field of characteristic 0 is again algebraically closed. The typeclass constraint \<^class>\field_gcd\ is a technical constraint that mandates that the field has a (trivial) GCD operation defined on it. It comes from some peculiarities of Isabelle's typeclass system and can be considered unimportant, since any concrete type of class \<^class>\field\ can easily be made an instance of \<^class>\field_gcd\. It would be possible to get rid of this constraint entirely here, but it is not worth the effort. The proof is a fairly standard one that uses Hensel's lemma. Some preliminary tricks are required to be able to use it, however, namely a number of non-obvious changes of variables to turn the polynomial with Puiseux coefficients into one with formal power series coefficients. The overall approach was taken from an article by Nowak~\<^cite>\"nowak2000"\. Basically, what we need to show is this: Let \[p(X,Z) = a_n(Z) X^n + a_{n-1}(Z) X^{n-1} + \ldots + a_0(Z)\] be a polynomial in \X\ of degree at least 2 with coefficients that are formal Puiseux series in \Z\. Then \p\ is reducible, i.e. it splits into two non-constant factors. Due to work we have already done elsewhere, we may assume here that $a_n = 1$, $a_{n-1} = 0$, and $a_0 \neq 0$, all of which will come in very useful. \ instance fpxs :: ("{alg_closed_field, field_char_0, field_gcd}") alg_closed_field proof (rule alg_closedI_reducible_coeff_deg_minus_one_eq_0) fix p :: "'a fpxs poly" assume deg_p: "degree p > 1" and lc_p: "lead_coeff p = 1" assume coeff_deg_minus_1: "coeff p (degree p - 1) = 0" assume "coeff p 0 \ 0" define N where "N = degree p" text \ Let $a_0, \ldots, a_n$ be the coefficients of \p\ with $a_n = 1$. Now let \r\ be the maximum of $-\frac{\text{val}(a_i)}{n-i}$ ranging over all $i < n$ such that $a_i \neq 0$. \ define r :: rat where "r = (MAX i\{i\{.. 0}. -fpxs_val (poly.coeff p i) / (rat_of_nat N - rat_of_nat i))" text \ We write $r = a / b$ such that all the $a_i$ can be written as Laurent series in $X^{1/b}$, i.e. the root orders of all the $a_i$ divide $b$: \ obtain a b where ab: "b > 0" "r = of_int a / of_nat b" "\i\N. fpxs_root_order (coeff p i) dvd b" proof - define b where "b = lcm (nat (snd (quotient_of r))) (LCM i\{..N}. fpxs_root_order (coeff p i))" define x where "x = b div nat (snd (quotient_of r))" define a where "a = fst (quotient_of r) * int x" show ?thesis proof (rule that) show "b > 0" using quotient_of_denom_pos'[of r] by (auto simp: b_def intro!: Nat.gr0I) have b_eq: "b = nat (snd (quotient_of r)) * x" by (simp add: x_def b_def) have "x > 0" using b_eq \b > 0\ by (auto intro!: Nat.gr0I) have "r = rat_of_int (fst (quotient_of r)) / rat_of_int (int (nat (snd (quotient_of r))))" using quotient_of_denom_pos'[of r] quotient_of_div[of r] by simp also have "\ = rat_of_int a / rat_of_nat b" using \x > 0\ by (simp add: a_def b_eq) finally show "r = rat_of_int a / rat_of_nat b" . show "\i\N. fpxs_root_order (poly.coeff p i) dvd b" by (auto simp: b_def) qed qed text \ We write all the coefficients of \p\ as Laurent series in $X^{1/b}$: \ have "\c. coeff p i = fpxs_compose_power (fpxs_of_fls c) (1 / rat_of_nat b)" if i: "i \ N" for i proof - have "fpxs_root_order (coeff p i) dvd b" using ab(3) i by auto from fpxs_as_fls'[OF this \b > 0\] show ?thesis by metis qed then obtain c_aux where c_aux: "coeff p i = fpxs_compose_power (fpxs_of_fls (c_aux i)) (1 / rat_of_nat b)" if "i \ N" for i by metis define c where "c = (\i. if i \ N then c_aux i else 0)" have c: "coeff p i = fpxs_compose_power (fpxs_of_fls (c i)) (1 / rat_of_nat b)" for i using c_aux[of i] by (auto simp: c_def N_def coeff_eq_0) have c_eq_0 [simp]: "c i = 0" if "i > N" for i using that by (auto simp: c_def) have c_eq: "fpxs_of_fls (c i) = fpxs_compose_power (coeff p i) (rat_of_nat b)" for i using c[of i] \b > 0\ by (simp add: fpxs_compose_power_distrib) text \ We perform another change of variables and multiply with a suitable power of \X\ to turn our Laurent coefficients into FPS coefficients: \ define c' where "c' = (\i. fls_X_intpow ((int N - int i) * a) * c i)" have "c' N = 1" using c[of N] \lead_coeff p = 1\ \b > 0\ by (simp add: c'_def N_def) have subdegree_c: "of_int (fls_subdegree (c i)) = fpxs_val (coeff p i) * rat_of_nat b" if i: "i \ N" for i proof - have "rat_of_int (fls_subdegree (c i)) = fpxs_val (fpxs_of_fls (c i))" by simp also have "fpxs_of_fls (c i) = fpxs_compose_power (poly.coeff p i) (rat_of_nat b)" by (subst c_eq) auto also have "fpxs_val \ = fpxs_val (coeff p i) * rat_of_nat b" using \b > 0\ by simp finally show ?thesis . qed text \ We now write all the coefficients as FPSs: \ have "\c''. c' i = fps_to_fls c''" if "i \ N" for i proof (cases "i = N") case True hence "c' i = fps_to_fls 1" using \c' N = 1\ by simp thus ?thesis by metis next case i: False show ?thesis proof (cases "c i = 0") case True hence "c' i = 0" by (auto simp: c'_def) - thus ?thesis by auto + thus ?thesis + by (metis fps_zero_to_fls) next case False hence "coeff p i \ 0" using c_eq[of i] by auto hence r_ge: "r \ -fpxs_val (poly.coeff p i) / (rat_of_nat N - rat_of_nat i)" unfolding r_def using i that False by (intro Max.coboundedI) auto have "fls_subdegree (c' i) = fls_subdegree (c i) + (int N - int i) * a" using i that False by (simp add: c'_def fls_X_intpow_times_conv_shift subdegree_c) also have "rat_of_int \ = fpxs_val (poly.coeff p i) * of_nat b + (of_nat N - of_nat i) * of_int a" using i that False by (simp add: subdegree_c) also have "\ = of_nat b * (of_nat N - of_nat i) * (fpxs_val (poly.coeff p i) / (of_nat N - of_nat i) + r)" using \b > 0\ i by (auto simp: field_simps ab(2)) also have "\ \ 0" using r_ge that by (intro mult_nonneg_nonneg) auto finally have "fls_subdegree (c' i) \ 0" by simp hence "\c''. c' i = fls_shift 0 (fps_to_fls c'')" by (intro fls_as_fps') (auto simp: algebra_simps) thus ?thesis by simp qed qed then obtain c''_aux where c''_aux: "c' i = fps_to_fls (c''_aux i)" if "i \ N" for i by metis define c'' where "c'' = (\i. if i \ N then c''_aux i else 0)" have c': "c' i = fps_to_fls (c'' i)" for i proof (cases "i \ N") case False thus ?thesis by (auto simp: c'_def c''_def) qed (auto simp: c''_def c''_aux) have c''_eq: "fps_to_fls (c'' i) = c' i" for i using c'[of i] by simp define p' where "p' = Abs_poly c''" have coeff_p': "coeff p' = c''" unfolding p'_def proof (rule coeff_Abs_poly) fix i assume "i > N" hence "coeff p i = 0" by (simp add: N_def coeff_eq_0) - thus "c'' i = 0" using c'[of i] c[of i] \b > 0\ - by (auto simp: c'_def fls_shift_eq0_iff) + thus "c'' i = 0" using c'[of i] c[of i] \b > 0\ \N < i\ c''_def by auto qed text \ We set up some homomorphisms to convert between the two polynomials: \ interpret comppow: map_poly_inj_idom_hom "(\x::'a fpxs. fpxs_compose_power x (1/rat_of_nat b))" by unfold_locales (use \b > 0\ in simp_all) define lift_poly :: "'a fps poly \ 'a fpxs poly" where "lift_poly = (\p. pcompose p [:0, fpxs_X_power r:]) \ (map_poly ((\x. fpxs_compose_power x (1/rat_of_nat b)) \ fpxs_of_fls \ fps_to_fls))" have [simp]: "degree (lift_poly q) = degree q" for q unfolding lift_poly_def by (simp add: degree_map_poly) interpret fps_to_fls: map_poly_inj_idom_hom fps_to_fls by unfold_locales (simp_all add: fls_times_fps_to_fls) interpret fpxs_of_fls: map_poly_inj_idom_hom fpxs_of_fls by unfold_locales simp_all interpret lift_poly: inj_idom_hom lift_poly unfolding lift_poly_def by (intro inj_idom_hom_compose inj_idom_hom_pcompose inj_idom_hom.inj_idom_hom_map_poly fps_to_fls.base.inj_idom_hom_axioms fpxs_of_fls.base.inj_idom_hom_axioms comppow.base.inj_idom_hom_axioms) simp_all interpret lift_poly: map_poly_inj_idom_hom lift_poly by unfold_locales define C :: "'a fpxs" where "C = fpxs_X_power (- (rat_of_nat N * r))" have [simp]: "C \ 0" by (auto simp: C_def) text \ Now, finally: the original polynomial and the new polynomial are related through the \<^term>\lift_poly\ homomorphism: \ have p_eq: "p = smult C (lift_poly p')" using \b > 0\ by (intro poly_eqI) (simp_all add: coeff_map_poly coeff_pcompose_linear coeff_p' c c''_eq c'_def C_def ring_distribs fpxs_X_power_conv_shift fpxs_shift_mult lift_poly_def ab(2) flip: fpxs_X_power_add fpxs_X_power_mult fpxs_shift_add_left) have [simp]: "degree p' = N" unfolding N_def using \b > 0\ by (simp add: p_eq) have lc_p': "lead_coeff p' = 1" using c''_eq[of N] by (simp add: coeff_p' \c' N = 1\) have "coeff p' (N - 1) = 0" using coeff_deg_minus_1 \b > 0\ unfolding N_def [symmetric] by (simp add: p_eq lift_poly_def coeff_map_poly coeff_pcompose_linear) text \ We reduce $p'(X,Z)$ to $p'(X,0)$: \ define p'_proj where "p'_proj = reduce_fps_poly p'" have [simp]: "degree p'_proj = N" unfolding p'_proj_def using lc_p' by (subst degree_reduce_fps_poly_monic) simp_all have lc_p'_proj: "lead_coeff p'_proj = 1" unfolding p'_proj_def using lc_p' by (subst reduce_fps_poly_monic) simp_all hence [simp]: "p'_proj \ 0" by auto have "coeff p'_proj (N - 1) = 0" using \coeff p' (N - 1) = 0\ by (simp add: p'_proj_def reduce_fps_poly_def) text \ We now show that \<^term>\p'_proj\ splits into non-trivial coprime factors. To do this, we have to show that it has two distinct roots, i.e. that it is not of the form $(X - c)^n$. \ obtain g h where gh: "degree g > 0" "degree h > 0" "coprime g h" "p'_proj = g * h" proof - have "degree p'_proj > 1" using deg_p by (auto simp: N_def) text \Let \x\ be an arbitrary root of \<^term>\p'_proj\:\ then obtain x where x: "poly p'_proj x = 0" using alg_closed_imp_poly_has_root[of p'_proj] by force text \Assume for the sake of contradiction that \<^term>\p'_proj\ were equal to $(1-x)^n$:\ have not_only_one_root: "p'_proj \ [:-x, 1:] ^ N" proof safe assume *: "p'_proj = [:-x, 1:] ^ N" text \ If \x\ were non-zero, all the coefficients of \p'_proj\ would also be non-zero by the Binomial Theorem. Since we know that the coefficient of \n - 1\ \<^emph>\is\ zero, this means that \x\ must be zero: \ have "coeff p'_proj (N - 1) = 0" by fact hence "x = 0" by (subst (asm) *, subst (asm) coeff_linear_poly_power) auto text \ However, by our choice of \r\, we know that there is an index \i\ such that \c' i\ has is non-zero and has valuation (i.e. subdegree) 0, which means that the \i\-th coefficient of \<^term>\p'_proj\ must also be non-zero. \ have "0 < N \ coeff p 0 \ 0" using deg_p \coeff p 0 \ 0\ by (auto simp: N_def) hence "{i\{.. 0} \ {}" by blast hence "r \ (\i. -fpxs_val (poly.coeff p i) / (rat_of_nat N - rat_of_nat i)) ` {i\{.. 0}" unfolding r_def using deg_p by (intro Max_in) (auto simp: N_def) then obtain i where i: "i < N" "coeff p i \ 0" "-fpxs_val (coeff p i) / (rat_of_nat N - rat_of_nat i) = r" by blast hence [simp]: "c' i \ 0" using i c[of i] by (auto simp: c'_def) have "fpxs_val (poly.coeff p i) = rat_of_int (fls_subdegree (c i)) / rat_of_nat b" using subdegree_c[of i] i \b > 0\ by (simp add: field_simps) also have "fpxs_val (coeff p i) = -r * (rat_of_nat N - rat_of_nat i)" using i by (simp add: field_simps) finally have "rat_of_int (fls_subdegree (c i)) = - r * (of_nat N - of_nat i) * of_nat b" using \b > 0\ by (simp add: field_simps) also have "c i = fls_shift ((int N - int i) * a) (c' i)" using i by (simp add: c'_def ring_distribs fls_X_intpow_times_conv_shift flip: fls_shifted_times_simps(2)) also have "fls_subdegree \ = fls_subdegree (c' i) - (int N - int i) * a" by (subst fls_shift_subdegree) auto finally have "fls_subdegree (c' i) = 0" using \b > 0\ by (simp add: ab(2)) hence "subdegree (coeff p' i) = 0" by (simp flip: c''_eq add: fls_subdegree_fls_to_fps coeff_p') moreover have "coeff p' i \ 0" using \c' i \ 0\ c' coeff_p' by auto ultimately have "coeff p' i $ 0 \ 0" using subdegree_eq_0_iff by blast also have "coeff p' i $ 0 = coeff p'_proj i" by (simp add: p'_proj_def reduce_fps_poly_def) also have "\ = 0" by (subst *, subst coeff_linear_poly_power) (use i \x = 0\ in auto) finally show False by simp qed text \ We can thus obtain our second root \y\ from the factorisation: \ have "\y. x \ y \ poly p'_proj y = 0" proof (rule ccontr) assume *: "\(\y. x \ y \ poly p'_proj y = 0)" have "p'_proj \ 0" by simp then obtain A where A: "size A = degree p'_proj" "p'_proj = smult (lead_coeff p'_proj) (\x\#A. [:-x, 1:])" using alg_closed_imp_factorization[of p'_proj] by blast have "set_mset A = {x. poly p'_proj x = 0}" using lc_p'_proj by (subst A) (auto simp: poly_prod_mset) also have "\ = {x}" using x * by auto finally have "A = replicate_mset N x" using set_mset_subset_singletonD[of A x] A(1) by simp with A(2) have "p'_proj = [:- x, 1:] ^ N" using lc_p'_proj by simp with not_only_one_root show False by contradiction qed then obtain y where "x \ y" "poly p'_proj y = 0" by blast text \ It now follows easily that \<^term>\p'_proj\ splits into non-trivial and coprime factors: \ show ?thesis proof (rule alg_closed_imp_poly_splits_coprime) show "degree p'_proj > 1" using deg_p by (simp add: N_def) show "x \ y" "poly p'_proj x = 0" "poly p'_proj y = 0" by fact+ qed (use that in metis) qed text \ By Hensel's lemma, these factors give rise to corresponding factors of \p'\: \ interpret hensel: fps_hensel p' p'_proj g h proof unfold_locales show "lead_coeff p' = 1" using lc_p' by simp qed (use gh \coprime g h\ in \simp_all add: p'_proj_def\) text \All that remains now is to undo the variable substitutions we did above:\ have "p = [:C:] * lift_poly hensel.G * lift_poly hensel.H" unfolding p_eq by (subst hensel.F_splits) (simp add: hom_distribs) thus "\irreducible p" by (rule reducible_polyI) (use hensel.deg_G hensel.deg_H gh in simp_all) qed text \ We do not actually show that this is the algebraic closure since this cannot be stated idiomatically in the typeclass setting and is probably not very useful either, but it can be motivated like this: Suppose we have an algebraically closed extension $L$ of the field of Laurent series. Clearly, $X^{a/b}\in L$ for any integer \a\ and any positive integer \b\ since $(X^{a/b})^b - X^a = 0$. But any Puiseux series $F(X)$ with root order \b\ can be written as \[F(X) = \sum_{k=0}^{b-1} X^{k/b} F_k(X)\] where the Laurent series $F_k(X)$ are defined as follows: \[F_k(X) := \sum_{n = n_{0,k}}^\infty [X^{n + k/b}] F(X) X^n\] Thus, $F(X)$ can be written as a finite sum of products of elements in $L$ and must therefore also be in $L$. Thus, the Puiseux series are all contained in $L$. \ subsection \Metric and topology\ text \ Formal Puiseux series form a metric space with the usual metric for formal series: Two series are ``close'' to one another if they have many initial coefficients in common. \ instantiation fpxs :: (zero) norm begin definition norm_fpxs :: "'a fpxs \ real" where "norm f = (if f = 0 then 0 else 2 powr (-of_rat (fpxs_val f)))" instance .. end instantiation fpxs :: (group_add) dist begin definition dist_fpxs :: "'a fpxs \ 'a fpxs \ real" where "dist f g = (if f = g then 0 else 2 powr (-of_rat (fpxs_val (f - g))))" instance .. end instantiation fpxs :: (group_add) metric_space begin definition uniformity_fpxs_def [code del]: "(uniformity :: ('a fpxs \ 'a fpxs) filter) = (INF e\{0 <..}. principal {(x, y). dist x y < e})" definition open_fpxs_def [code del]: "open (U :: 'a fpxs set) \ (\x\U. eventually (\(x', y). x' = x \ y \ U) uniformity)" instance proof fix f g h :: "'a fpxs" show "dist f g \ dist f h + dist g h" proof (cases "f \ g \ f \ h \ g \ h") case True have "dist f g \ 2 powr -real_of_rat (min (fpxs_val (f - h)) (fpxs_val (g - h)))" using fpxs_val_add_ge[of "f - h" "h - g"] True by (auto simp: algebra_simps fpxs_val_minus_commute dist_fpxs_def of_rat_less_eq) also have "\ \ dist f h + dist g h" using True by (simp add: dist_fpxs_def min_def) finally show ?thesis . qed (auto simp: dist_fpxs_def fpxs_val_minus_commute) qed (simp_all add: uniformity_fpxs_def open_fpxs_def dist_fpxs_def) end instance fpxs :: (group_add) dist_norm by standard (auto simp: dist_fpxs_def norm_fpxs_def) end \ No newline at end of file diff --git a/thys/Formal_Puiseux_Series/Puiseux_Laurent_Library.thy b/thys/Formal_Puiseux_Series/Puiseux_Laurent_Library.thy deleted file mode 100644 --- a/thys/Formal_Puiseux_Series/Puiseux_Laurent_Library.thy +++ /dev/null @@ -1,212 +0,0 @@ -(* - File: Puiseux_Laurent_Library.thy - Author: Manuel Eberl, TU München -*) -subsection \Facts about Laurent series\ -theory Puiseux_Laurent_Library - imports "HOL-Computational_Algebra.Computational_Algebra" -begin - -lemma filterlim_at_top_div_const_nat: - assumes "c > 0" - shows "filterlim (\x::nat. x div c) at_top at_top" - unfolding filterlim_at_top -proof - fix C :: nat - have *: "n div c \ C" if "n \ C * c" for n - using assms that by (metis div_le_mono div_mult_self_is_m) - have "eventually (\n. n \ C * c) at_top" - by (rule eventually_ge_at_top) - thus "eventually (\n. n div c \ C) at_top" - by eventually_elim (use * in auto) -qed - -lemma fls_eq_iff: "f = g \ (\n. fls_nth f n = fls_nth g n)" - by transfer auto - -lemma fls_shift_eq_1_iff: "fls_shift n f = 1 \ f = fls_shift (-n) 1" -proof - - have "fls_shift n f = 1 \ fls_shift n f = fls_shift n (fls_shift (-n) 1)" - by (simp del: fls_shift_eq_iff) - also have "\ \ f = fls_shift (-n) 1" - by (subst fls_shift_eq_iff) auto - finally show ?thesis . -qed - -lemma fps_to_fls_eq_iff [simp]: "fps_to_fls f = fps_to_fls g \ f = g" -proof safe - assume "fps_to_fls f = fps_to_fls g" - hence *: "n < 0 \ f $ nat n = g $ nat n" for n - by (force simp: fls_eq_iff split: if_splits) - have "f $ n = g $ n" for n - using *[of "int n"] by auto - thus "f = g" - by (auto simp: fps_eq_iff) -qed - -lemma fps_to_fls_eq_0_iff [simp]: "fps_to_fls f = 0 \ f = 0" - using fps_to_fls_eq_iff[of f 0] by (simp del: fps_to_fls_eq_iff) - -lemma fps_to_fls_eq_1_iff [simp]: "fps_to_fls f = 1 \ f = 1" - using fps_to_fls_eq_iff[of f 1] by (simp del: fps_to_fls_eq_iff) - -lemma fps_to_fls_power: "fps_to_fls (f ^ n) = fps_to_fls f ^ n" - by (induction n) (auto simp: fls_times_fps_to_fls) - -lemma fls_as_fps: - fixes f :: "'a :: zero fls" and n :: int - assumes n: "n \ -fls_subdegree f" - obtains f' where "f = fls_shift n (fps_to_fls f')" -proof - - have "fls_subdegree (fls_shift (- n) f) \ 0" - by (rule fls_shift_nonneg_subdegree) (use n in simp) - hence "f = fls_shift n (fps_to_fls (fls_regpart (fls_shift (-n) f)))" - by (subst fls_regpart_to_fls_trivial) simp_all - thus ?thesis - by (rule that) -qed - -lemma fls_as_fps': - fixes f :: "'a :: zero fls" and n :: int - assumes n: "n \ -fls_subdegree f" - shows "\f'. f = fls_shift n (fps_to_fls f')" - using fls_as_fps[OF assms] by metis - - -definition fls_compose_fps :: "'a :: field fls \ 'a fps \ 'a fls" where - "fls_compose_fps f g = - (if f = 0 then 0 - else if fls_subdegree f \ 0 then fps_to_fls (fps_compose (fls_regpart f) g) - else fps_to_fls (fps_compose (fls_base_factor_to_fps f) g) / - fps_to_fls g ^ nat (-fls_subdegree f))" - -lemma fls_compose_fps_fps [simp]: - "fls_compose_fps (fps_to_fls f) g = fps_to_fls (fps_compose f g)" - by (auto simp: fls_compose_fps_def fls_subdegree_fls_to_fps) - -lemma fls_const_transfer [transfer_rule]: - "rel_fun (=) (pcr_fls (=)) - (\c n. if n = 0 then c else 0) fls_const" - by (auto simp: fls_const_def rel_fun_def pcr_fls_def OO_def cr_fls_def) - -lemma fls_shift_transfer [transfer_rule]: - "rel_fun (=) (rel_fun (pcr_fls (=)) (pcr_fls (=))) - (\n f k. f (k+n)) fls_shift" - by (auto simp: fls_const_def rel_fun_def pcr_fls_def OO_def cr_fls_def) - -lift_definition fls_compose_power :: "'a :: zero fls \ nat \ 'a fls" is - "\f d n. if d > 0 \ int d dvd n then f (n div int d) else 0" -proof - - fix f :: "int \ 'a" and d :: nat - assume *: "eventually (\n. f (-int n) = 0) cofinite" - show "eventually (\n. (if d > 0 \ int d dvd -int n then f (-int n div int d) else 0) = 0) cofinite" - proof (cases "d = 0") - case False - from * have "eventually (\n. f (-int n) = 0) at_top" - by (simp add: cofinite_eq_sequentially) - hence "eventually (\n. f (-int (n div d)) = 0) at_top" - by (rule eventually_compose_filterlim[OF _ filterlim_at_top_div_const_nat]) (use False in auto) - hence "eventually (\n. (if d > 0 \ int d dvd -int n then f (-int n div int d) else 0) = 0) at_top" - by eventually_elim (auto simp: zdiv_int dvd_neg_div) - thus ?thesis - by (simp add: cofinite_eq_sequentially) - qed auto -qed - -lemma fls_nth_compose_power: - assumes "d > 0" - shows "fls_nth (fls_compose_power f d) n = (if int d dvd n then fls_nth f (n div int d) else 0)" - using assms by transfer auto - - -lemma fls_compose_power_0_left [simp]: "fls_compose_power 0 d = 0" - by transfer auto - -lemma fls_compose_power_1_left [simp]: "d > 0 \ fls_compose_power 1 d = 1" - by transfer (auto simp: fun_eq_iff) - -lemma fls_compose_power_const_left [simp]: - "d > 0 \ fls_compose_power (fls_const c) d = fls_const c" - by transfer (auto simp: fun_eq_iff) - -lemma fls_compose_power_shift [simp]: - "d > 0 \ fls_compose_power (fls_shift n f) d = fls_shift (d * n) (fls_compose_power f d)" - by transfer (auto simp: fun_eq_iff add_ac mult_ac) - -lemma fls_compose_power_X_intpow [simp]: - "d > 0 \ fls_compose_power (fls_X_intpow n) d = fls_X_intpow (int d * n)" - by simp - -lemma fls_compose_power_X [simp]: - "d > 0 \ fls_compose_power fls_X d = fls_X_intpow (int d)" - by transfer (auto simp: fun_eq_iff) - -lemma fls_compose_power_X_inv [simp]: - "d > 0 \ fls_compose_power fls_X_inv d = fls_X_intpow (-int d)" - by (simp add: fls_X_inv_conv_shift_1) - -lemma fls_compose_power_0_right [simp]: "fls_compose_power f 0 = 0" - by transfer auto - -lemma fls_compose_power_add [simp]: - "fls_compose_power (f + g) d = fls_compose_power f d + fls_compose_power g d" - by transfer auto - -lemma fls_compose_power_diff [simp]: - "fls_compose_power (f - g) d = fls_compose_power f d - fls_compose_power g d" - by transfer auto - -lemma fls_compose_power_uminus [simp]: - "fls_compose_power (-f) d = -fls_compose_power f d" - by transfer auto - -lemma fps_nth_compose_X_power: - "fps_nth (f oo (fps_X ^ d)) n = (if d dvd n then fps_nth f (n div d) else 0)" -proof - - have "fps_nth (f oo (fps_X ^ d)) n = (\i = 0..n. f $ i * (fps_X ^ (d * i)) $ n)" - unfolding fps_compose_def by (simp add: power_mult) - also have "\ = (\i\(if d dvd n then {n div d} else {}). f $ i * (fps_X ^ (d * i)) $ n)" - by (intro sum.mono_neutral_right) auto - also have "\ = (if d dvd n then fps_nth f (n div d) else 0)" - by auto - finally show ?thesis . -qed - -lemma fls_compose_power_fps_to_fls: - assumes "d > 0" - shows "fls_compose_power (fps_to_fls f) d = fps_to_fls (fps_compose f (fps_X ^ d))" - using assms - by (intro fls_eqI) (auto simp: fls_nth_compose_power fps_nth_compose_X_power - pos_imp_zdiv_neg_iff div_neg_pos_less0 nat_div_distrib - simp flip: int_dvd_int_iff) - -lemma fls_compose_power_mult [simp]: - "fls_compose_power (f * g :: 'a :: idom fls) d = fls_compose_power f d * fls_compose_power g d" -proof (cases "d > 0") - case True - define n where "n = nat (max 0 (max (- fls_subdegree f) (- fls_subdegree g)))" - have n_ge: "-fls_subdegree f \ int n" "-fls_subdegree g \ int n" - unfolding n_def by auto - obtain f' where f': "f = fls_shift n (fps_to_fls f')" - using fls_as_fps[OF n_ge(1)] by (auto simp: n_def) - obtain g' where g': "g = fls_shift n (fps_to_fls g')" - using fls_as_fps[OF n_ge(2)] by (auto simp: n_def) - show ?thesis using \d > 0\ - by (simp add: f' g' fls_shifted_times_simps mult_ac fls_compose_power_fps_to_fls - fps_compose_mult_distrib flip: fls_times_fps_to_fls) -qed auto - -lemma fls_compose_power_power [simp]: - assumes "d > 0 \ n > 0" - shows "fls_compose_power (f ^ n :: 'a :: idom fls) d = fls_compose_power f d ^ n" -proof (cases "d > 0") - case True - thus ?thesis by (induction n) auto -qed (use assms in auto) - -lemma fls_nth_compose_power' [simp]: - "d = 0 \ \d dvd n \ fls_nth (fls_compose_power f d) n = 0" - "d dvd n \ d > 0 \ fls_nth (fls_compose_power f d) n = fls_nth f (n div d)" - by (transfer; force; fail)+ - -end \ No newline at end of file diff --git a/thys/Formal_Puiseux_Series/ROOT b/thys/Formal_Puiseux_Series/ROOT --- a/thys/Formal_Puiseux_Series/ROOT +++ b/thys/Formal_Puiseux_Series/ROOT @@ -1,15 +1,14 @@ chapter AFP session Formal_Puiseux_Series (AFP) = "HOL-Computational_Algebra" + options [timeout = 1200] sessions Polynomial_Interpolation theories Puiseux_Polynomial_Library - Puiseux_Laurent_Library FPS_Hensel Formal_Puiseux_Series document_files "root.tex" "root.bib" diff --git a/thys/Hermite_Lindemann/Hermite_Lindemann.thy b/thys/Hermite_Lindemann/Hermite_Lindemann.thy --- a/thys/Hermite_Lindemann/Hermite_Lindemann.thy +++ b/thys/Hermite_Lindemann/Hermite_Lindemann.thy @@ -1,2558 +1,2561 @@ (* File: Hermite_Lindemann.thy Author: Manuel Eberl, TU München *) section \The Hermite--Lindemann--Weierstraß Transcendence Theorem\ theory Hermite_Lindemann imports Pi_Transcendental.Pi_Transcendental Algebraic_Numbers.Algebraic_Numbers Algebraic_Integer_Divisibility More_Min_Int_Poly Complex_Lexorder More_Polynomial_HLW More_Multivariate_Polynomial_HLW More_Algebraic_Numbers_HLW Misc_HLW begin +hide_const (open) Henstock_Kurzweil_Integration.content Module.smult + + text \ The Hermite--Lindemann--Weierstraß theorem answers questions about the transcendence of the exponential function and other related complex functions. It proves that a large number of combinations of exponentials is always transcendental. A first (much weaker) version of the theorem was proven by Hermite. Lindemann and Weierstraß then successively generalised it shortly afterwards, and finally Baker gave another, arguably more elegant formulation (which is the one that we will prove, and then derive the traditional version from it). To honour the contributions of all three of these 19th-century mathematicians, I refer to the theorem as the Hermite--Lindemann--Weierstraß theorem, even though in other literature it is often called Hermite--Lindemann or Lindemann--Weierstraß. To keep things short, the Isabelle name of the theorem, however, will omit Weierstraß's name. \ subsection \Main proof\ text \ Following Baker, We first prove the following special form of the theorem: Let $m > 0$ and $q_1, \ldots, q_m \in\mathbb{Z}[X]$ be irreducible, non-constant, and pairwise coprime polynomials. Let $\beta_1, \ldots, \beta_m$ be non-zero integers. Then \[\sum_{i=1}^m \beta_i \sum_{q_i(\alpha) = 0} e^\alpha \neq 0\] The difference to the final theorem is that \<^enum> The coefficients $\beta_i$ are non-zero integers (as opposed to arbitrary algebraic numbers) \<^enum> The exponents $\alpha_i$ occur in full sets of conjugates, and each set has the same coefficient. In a similar fashion to the proofs of the transcendence of \e\ and \\\, we define some number $J$ depending on the $\alpha_i$ and $\beta_i$ and an arbitrary sufficiently large prime \p\. We then show that, on one hand, $J$ is an integer multiple of $(p-1)!$, but on the other hand it is bounded from above by a term of the form $C_1 \cdot C_2^p$. This is then clearly a contradiction if \p\ is chosen large enough. \ lemma Hermite_Lindemann_aux1: fixes P :: "int poly set" and \ :: "int poly \ int" assumes "finite P" and "P \ {}" assumes distinct: "pairwise Rings.coprime P" assumes irred: "\p. p \ P \ irreducible p" assumes nonconstant: "\p. p \ P \ Polynomial.degree p > 0" assumes \_nz: "\p. p \ P \ \ p \ 0" defines "Roots \ (\p. {\::complex. poly (of_int_poly p) \ = 0})" shows "(\p\P. of_int (\ p) * (\\\Roots p. exp \)) \ 0" proof note [intro] = \finite P\ assume sum_eq_0: "(\p\P. of_int (\ p) * (\\\Roots p. exp \)) = 0" define Roots' where "Roots' = (\p\P. Roots p)" have finite_Roots [intro]: "finite (Roots p)" if "p \ P" for p using nonconstant[of p] that by (auto intro: poly_roots_finite simp: Roots_def) have [intro]: "finite Roots'" by (auto simp: Roots'_def) have [simp]: "0 \ P" using nonconstant[of 0] by auto have [simp]: "p \ 0" if "p \ P" for p using that by auto text \ The polynomials in \<^term>\P\ do not have multiple roots: \ have rsquarefree: "rsquarefree (of_int_poly q :: complex poly)" if "q \ P" for q by (rule irreducible_imp_rsquarefree_of_int_poly) (use that in \auto intro: irred nonconstant\) text \ No two different polynomials in \<^term>\P\ have roots in common: \ have disjoint: "disjoint_family_on Roots P" using distinct proof (rule pairwise_imp_disjoint_family_on) fix p q assume P: "p \ P" "q \ P" and "Rings.coprime p q" hence "Rings.coprime (of_int_poly p :: complex poly) (of_int_poly q)" by (intro coprime_of_int_polyI) thus "Roots p \ Roots q = {}" using poly_eq_0_coprime[of "of_int_poly p" "of_int_poly q :: complex poly"] P by (auto simp: Roots_def) qed define n_roots :: "int poly \ nat" ("\_") where "n_roots = (\p. card (Roots p))" define n where "n = (\p\P. \p)" have n_altdef: "n = card Roots'" unfolding n_def Roots'_def n_roots_def using disjoint by (subst card_UN_disjoint) (auto simp: disjoint_family_on_def) have Roots_nonempty: "Roots p \ {}" if "p \ P" for p using nonconstant[OF that] by (auto simp: Roots_def fundamental_theorem_of_algebra constant_degree) have "Roots' \ {}" using Roots_nonempty \P \ {}\ by (auto simp: Roots'_def) have "n > 0" using \Roots' \ {}\ \finite Roots'\ by (auto simp: n_altdef) text \ We can split each polynomial in \P\ into a product of linear factors: \ have of_int_poly_P: "of_int_poly q = Polynomial.smult (Polynomial.lead_coeff q) (\x\Roots q. [:-x, 1:])" if "q \ P" for q using complex_poly_decompose_rsquarefree[OF rsquarefree[OF that]] by (simp add: Roots_def) text \ We let \l\ be an integer such that $l\alpha$ is an algebraic integer for all our roots \\\: \ define l where "l = (LCM q\P. Polynomial.lead_coeff q)" have alg_int: "algebraic_int (of_int l * x)" if "x \ Roots'" for x proof - from that obtain q where q: "q \ P" "ipoly q x = 0" by (auto simp: Roots'_def Roots_def) show ?thesis by (rule algebraic_imp_algebraic_int'[of q]) (use q in \auto simp: l_def\) qed have "l \ 0" using \finite P\ by (auto simp: l_def Lcm_0_iff) moreover have "l \ 0" unfolding l_def by (rule Lcm_int_greater_eq_0) ultimately have "l > 0" by linarith text \ We can split the product of all the polynomials in \P\ into linear factors: \ define lc_factor where "lc_factor = (\q\P. l ^ Polynomial.degree q div Polynomial.lead_coeff q)" have lc_factor: "Polynomial.smult (of_int l ^ n) (\\'\Roots'. [:-\',1:]) = of_int_poly (Polynomial.smult lc_factor (\P))" proof - define lc where "lc = (\q. Polynomial.lead_coeff q :: int)" define d where "d = (Polynomial.degree :: int poly \ nat)" have "(\q\P. of_int_poly q) = (\q\P. Polynomial.smult (lc q) (\x\Roots q. [:-x, 1:]) :: complex poly)" unfolding lc_def by (intro prod.cong of_int_poly_P refl) also have "\ = Polynomial.smult (\q\P. lc q) (\q\P. (\x\Roots q. [:-x, 1:]))" by (simp add: prod_smult) also have "(\q\P. (\x\Roots q. [:-x, 1:])) = (\x\Roots'. [:-x, 1:])" unfolding Roots'_def using disjoint by (intro prod.UNION_disjoint [symmetric]) (auto simp: disjoint_family_on_def) also have "Polynomial.smult (of_int lc_factor) (Polynomial.smult (\q\P. lc q) \) = Polynomial.smult (\q\P. of_int (l ^ d q div lc q * lc q)) (\x\Roots'. pCons (- x) 1)" by (simp add: lc_factor_def prod.distrib lc_def d_def) also have "(\q\P. of_int (l ^ d q div lc q * lc q)) = (\q\P. of_int l ^ d q :: complex)" proof (intro prod.cong, goal_cases) case (2 q) have "lc q dvd l" unfolding l_def lc_def using 2 by auto also have "\ dvd l ^ d q" using 2 nonconstant[of q] by (intro dvd_power) (auto simp: d_def) finally show ?case by simp qed auto also have "\ = l ^ (\q\P. d q)" by (simp add: power_sum) also have "(\q\P. d q) = (\q\P. n_roots q)" proof (intro sum.cong, goal_cases) case (2 q) thus ?case using rsquarefree[OF 2] by (subst (asm) rsquarefree_card_degree) (auto simp: d_def n_roots_def Roots_def) qed auto also have "\ = n" by (simp add: n_def) finally show ?thesis by (simp add: of_int_hom.map_poly_hom_smult of_int_poly_hom.hom_prod) qed text \ We define \R\ to be the radius of the smallest circle around the origin in which all our roots lie: \ define R :: real where "R = Max (norm ` Roots')" have R_ge: "R \ norm \" if "\ \ Roots'" for \ unfolding R_def using that by (intro Max_ge) auto have "R \ 0" proof - from \Roots' \ {}\ obtain \ where "\ \ Roots'" by auto have "0 \ norm \" by simp also have "\ \ R" by (intro R_ge) fact finally show "R \ 0" by simp qed text \ Now the main part of the proof: for any sufficiently large prime \p\, our assumptions imply $(p-1)! ^ n \leq C' l^{np} (2R)^{np-1}$ for some constant $C'$: \ define C :: "nat \ real" where "C = (\p. l ^ (n * p) * (2*R) ^ (n * p - 1))" define C' where "C' = (\x\Roots'. \q\P. real_of_int \\ q\ * (\\\Roots q. cmod \ * exp (cmod \)))" text \ We commence with the proof of the main inequality. \ have ineq: "fact (p - 1) ^ n \ C' * C p ^ n" if p: "prime p" and p_ineqs: "\q\P. p > \\ q\" "real p > norm (\\\Roots'. of_int (l^n) * (\x\Roots'-{\}. \ - x))" for p :: nat proof - have "p > 1" using prime_gt_1_nat[OF p] . text \ We define the polynomial function \[f_i(X) = l^{np} \frac{\prod_\alpha (X-\alpha)^p}{X - \alpha_i}\] where the product runs over all roots $\alpha$. \ define f_poly :: "complex \ complex poly" where "f_poly = (\\. Polynomial.smult (l^(n*p)) ((\\'\Roots'. [:-\', 1:]^p) div [:-\, 1:]))" have f_poly_altdef: "f_poly \ = Polynomial.smult (l^(n*p)) ((\\'\Roots'. [:-\', 1:]^(if \' = \ then p - 1 else p)))" if "\ \ Roots'" for \ proof - have "(\\'\Roots'. [:-\', 1:] ^ (if \'=\ then p-1 else p)) * [:-\, 1:] = [:- \, 1:] ^ (p - 1) * (\x\Roots' - {\}. [:- x, 1:] ^ p) * [:- \, 1:]" using that by (subst prod.If_eq) (auto simp: algebra_simps) also have "\ = (\x\Roots' - {\}. [:- x, 1:] ^ p) * [:- \, 1:] ^ Suc (p - 1)" by (simp only: power_Suc mult_ac) also have "Suc (p - 1) = p" using \p > 1\ by auto also have "(\x\Roots' - {\}. [:- x, 1:] ^ p) * [:- \, 1:] ^ p = (\x\Roots'. [:- x, 1:] ^ p)" using that by (subst prod.remove[of _ \]) auto finally have eq: "(\\'\Roots'. [:-\', 1:] ^ (if \'=\ then p-1 else p)) * [:-\, 1:] = (\x\Roots'. [:- x, 1:] ^ p)" . show ?thesis unfolding f_poly_def eq[symmetric] by (subst nonzero_mult_div_cancel_right) auto qed define f :: "complex \ complex \ complex" where "f = (\\ x. l^(n*p) * (\\'\Roots'. (x - \')^(if \' = \ then p - 1 else p)))" have eval_f: "poly (f_poly \) x = f \ x" if "\ \ Roots'" for \ x using that by (simp add: f_poly_altdef poly_prod f_def) have deg_f: "Polynomial.degree (f_poly \) = n * p - 1" if "\ \ Roots'" for \ proof - have "Polynomial.degree (f_poly \) = p - 1 + (n - 1) * p" unfolding f_poly_altdef[OF that] using that \l > 0\ \finite Roots'\ by (subst prod.If_eq) (auto simp: degree_prod_eq degree_power_eq degree_mult_eq n_altdef) also have "p - 1 + (n - 1) * p = n * p - 1" using \n > 0\ \p > 1\ by (cases n) auto finally show ?thesis . qed text \ Next, we define the function $I_i(z) = \int_0^z e^{z-t} f_i(t) \text{d}t$, and, based on that, the numbers $J_i = \sum_{i=1}^m \beta_i \sum_{q_i(x) = 0} I_i(x)$, and the number $J$, which is the product of all the $J_i$: \ define I :: "complex \ complex \ complex" where "I = (\\ x. lindemann_weierstrass_aux.I (f_poly \) x)" define J :: "complex \ complex" where "J = (\\. \q\P. \ q * (\x\Roots q. I \ x))" define J' :: complex where "J' = (\\\Roots'. J \)" text \ Reusing some of the machinery from the proof that \e\ is transcendental, we find the following equality for $J_i$: \ have J_eq: "J \ = -(\q\P. of_int (\ q) * (\x\Roots q. \j)) x))" if "\ \ Roots'" for \ proof - have "n * p \ 1 * 2" using \n > 0\ \p > 1\ by (intro mult_mono) auto hence [simp]: "{..n*p-Suc 0} = {.. = (\q\P. \ q * (\x\Roots q. I \ x))" unfolding J_def .. also have "\ = (\q\P. of_int (\ q) * (\x\Roots q. exp x * (\j)) 0))) - (\q\P. of_int (\ q) * (\x\Roots q. \j)) x))" unfolding I_def lindemann_weierstrass_aux.I_def by (simp add: deg_f that ring_distribs sum_subtractf sum_distrib_left sum_distrib_right mult_ac) also have "\ = -(\q\P. of_int (\ q) * (\x\Roots q. \j)) x))" unfolding sum_distrib_right [symmetric] mult.assoc [symmetric] sum_eq_0 by simp finally show ?thesis . qed text \ The next big step is to show that $(p-1)! \mid J_i$ as an algebraic integer (i.e. $J_i / (p-1)!$ is an algebraic integer), but $p \not\mid J_i$. This is done by brute force: We show that every summand in the above sum has $p!$ as a factor, except for the one corresponding to $x = \alpha_i$, $j = p - 1$, which has $(p-1)!$ as a factor but not \p\. \ have J: "fact (p - 1) alg_dvd J \" "\of_nat p alg_dvd J \" if \: "\ \ Roots'" for \ proof - define h where "h = (\\' j. poly ((pderiv ^^ j) (f_poly \)) \')" from \ obtain q where q: "q \ P" "\ \ Roots q" by (auto simp: Roots'_def) have "J \ = -(\(q, \')\Sigma P Roots. \j q) * h \' j)" unfolding J_eq[OF \] h_def sum_distrib_left by (subst (2) sum.Sigma) auto also have "\ = -(\((q,\'),i)\Sigma P Roots\{.. q) * h \' i)" by (subst (2) sum.Sigma [symmetric]) (auto simp: case_prod_unfold) finally have J_eq': "J \ = - (\((q, \'), i)\Sigma P Roots \ {.. q) * h \' i)" . have h_\_pm1_eq: "h \ (p-1) = of_int (l^(n*p)) * fact (p-1) * (\\'\Roots'-{\}. (\-\')^p)" proof - have "h \ (p - 1) = of_int (l ^ (n * p)) * poly ((pderiv ^^ (p-1)) (\\'\Roots'. [:-\',1:] ^ (if \' = \ then p - 1 else p))) \" unfolding h_def f_poly_altdef[OF \] higher_pderiv_smult poly_smult .. also have "(\\'\Roots'. [:-\',1:] ^ (if \' = \ then p - 1 else p)) = [:-\,1:]^(p-1) * (\\'\Roots'-{\}. [:-\',1:]^p)" using \ by (subst prod.If_eq) auto also have "poly ((pderiv ^^ (p-1)) \) \ = fact (p - 1) * (\\'\Roots' - {\}. (\ - \') ^ p)" by (subst poly_higher_pderiv_aux2) (simp_all add: poly_prod) finally show ?thesis by (simp only: mult.assoc) qed have "fact (p-1) alg_dvd h \ (p-1)" proof - have "fact (p-1) alg_dvd fact (p-1) * (of_int (l^p) * (\\'\Roots'-{\}. (l*\-l*\')^p))" by (intro alg_dvd_triv_left algebraic_int_times[of "of_int (l^p)"] algebraic_int_prod algebraic_int_power algebraic_int_diff alg_int \ algebraic_int_of_int) auto also have "(\\'\Roots'-{\}. (l*\-l*\')^p) = (\\'\Roots'-{\}. of_int l^p * (\-\')^p)" by (subst power_mult_distrib [symmetric]) (simp_all add: algebra_simps) also have "\ = of_int (l ^ (p * (n-1))) * (\\'\Roots'-{\}. (\-\')^p)" using \ by (subst prod.distrib) (auto simp: card_Diff_subset n_altdef simp flip: power_mult) also have "of_int (l^p) * \ = of_int (l^(p+p*(n-1))) * (\\'\Roots'-{\}. (\-\')^p)" unfolding mult.assoc [symmetric] power_add [symmetric] of_int_power .. also have "p + p * (n - 1) = n * p" using \n > 0\ by (cases n) (auto simp: mult_ac) also have "fact (p - 1) * (of_int (l^(n*p)) * (\\'\Roots'-{\}. (\-\')^p)) = h \ (p-1)" unfolding h_\_pm1_eq by (simp add: mult_ac) finally show ?thesis . qed have "\of_nat p alg_dvd of_int (\ q) * h \ (p-1)" unfolding h_\_pm1_eq mult.assoc [symmetric] of_int_mult [symmetric] proof define r where "r = (\\. of_int (l^n) * (\\'\Roots'-{\}. \-\'))" have alg_int_r: "algebraic_int (r \)" if "\ \ Roots'" for \ proof - have "algebraic_int (of_int l * (\\'\Roots'-{\}. of_int l * \ - of_int l * \'))" by (intro algebraic_int_times[OF algebraic_int_of_int] algebraic_int_prod algebraic_int_power algebraic_int_diff alg_int that) auto also have "\ = of_int l * (\\'\Roots'-{\}. of_int l * (\ - \'))" by (simp add: algebra_simps flip: power_mult_distrib) also have "\ = of_int (l^(1 + (n-1))) * (\\'\Roots'-{\}. \ - \')" using that by (simp add: r_def prod.distrib card_Diff_subset n_altdef power_add mult_ac flip: power_mult) also have "1 + (n - 1) = n" using \n > 0\ by auto finally show "algebraic_int (r \)" unfolding r_def . qed have "(\\'\Roots'. r \') \ \" proof - obtain Root where Root_bij: "bij_betw Root {..finite Roots'\] unfolding n_altdef atLeast0LessThan by metis have Root_in_Roots': "Root i \ Roots'" if "i < n" for i using Root_bij that by (auto simp: bij_betw_def) define R :: "complex mpoly" where "R = (\ij\{.. \" proof (rule symmetric_poly_of_roots_in_subring) show "symmetric_mpoly {..\. \"], goal_cases) case (2 i \) from \\ permutes {.. have [simp]: "bij \" by (rule permutes_bij) have "mpoly_map_vars \ (Const (of_int (l ^ n)) * (\j\{..j\{.. i) - Var (\ j))" by simp also have "(\j\{.. i) - Var (\ j)) = (\j\{.. i}. Var (\ i) - Var j)" using 2 permutes_in_image[OF 2(2), of i] by (intro prod.reindex_bij_betw bij_betw_Diff permutes_imp_bij[OF 2(2)]) (auto simp: bij_betw_singleton) finally show ?case by simp qed next show "vars R \ {.. :: complex set)" by unfold_locales auto then interpret ring_closed "\ :: complex set" . show "\m. MPoly_Type.coeff R m \ \" unfolding R_def by (intro allI coeff_prod_closed coeff_mult_closed coeff_power_closed) (auto simp: mpoly_coeff_Const coeff_Var when_def) next let ?lc = "of_int (\p\P. Polynomial.lead_coeff p) :: complex" have "(\q\P. of_int_poly q) = (\q\P. Polynomial.smult (of_int (Polynomial.lead_coeff q)) (\x\Roots q. [:-x, 1:]))" by (intro prod.cong of_int_poly_P refl) also have "\ = Polynomial.smult ?lc (\q\P. \x\Roots q. [:-x, 1:])" by (simp add: prod_smult) also have "(\q\P. \x\Roots q. [:-x, 1:]) = (\x\Roots'. [:-x, 1:])" unfolding Roots'_def using disjoint by (intro prod.UNION_disjoint [symmetric]) (auto simp: disjoint_family_on_def) also have "\ = (\iP) = Polynomial.smult ?lc (\i 0" by (intro prod_nonzeroI) auto thus "inverse ?lc * ?lc = 1" "inverse ?lc \ \" by (auto simp: field_simps simp flip: of_int_prod) qed auto also have "insertion Root R = (\ij\{.. = (\i\'\Roots'-{Root i}. Root i - \'))" proof (intro prod.cong, goal_cases) case (2 i) hence "(\j\{..\'\Roots'-{Root i}. Root i - \')" by (intro prod.reindex_bij_betw bij_betw_Diff Root_bij) (auto intro: Root_in_Roots' simp: bij_betw_singleton) thus ?case by simp qed auto also have "\ = (\\'\Roots'. r \')" unfolding r_def by (intro prod.reindex_bij_betw Root_bij) finally show "(\\'\Roots'. r \') \ \" . qed moreover have "algebraic_int (\\'\Roots'. r \')" by (intro algebraic_int_prod alg_int_r) ultimately have is_int: "(\\'\Roots'. r \') \ \" using rational_algebraic_int_is_int by blast then obtain R' where R': "(\\'\Roots'. r \') = of_int R'" by (elim Ints_cases) have "(\\'\Roots'. r \') \ 0" using \l > 0\ by (intro prod_nonzeroI) (auto simp: r_def \finite Roots'\) with R' have [simp]: "R' \ 0" by auto assume "of_nat p alg_dvd of_int (\ q * l^(n*p)) * fact (p-1) * (\\'\Roots'-{\}. (\-\') ^ p)" also have "\ = of_int (\ q) * fact (p-1) * r \ ^ p" by (simp add: r_def mult_ac power_mult_distrib power_mult prod_power_distrib) also have "\ alg_dvd of_int (\ q) * fact (p-1) * r \ ^ p * (\\'\Roots'-{\}. r \') ^ p" by (intro alg_dvd_triv_left algebraic_int_prod alg_int_r algebraic_int_power) auto also have "\ = of_int (\ q) * fact (p-1) * (\\'\Roots'. r \') ^ p" using \ by (subst (2) prod.remove[of _ "\"]) (auto simp: mult_ac power_mult_distrib) also have "\ = of_int (\ q * fact (p - 1) * R' ^ p)" by (simp add: R') also have "of_nat p = of_int (int p)" by simp finally have "int p dvd \ q * fact (p - 1) * R' ^ p" by (subst (asm) alg_dvd_of_int_iff) moreover have "prime (int p)" using \prime p\ by auto ultimately have "int p dvd \ q \ int p dvd fact (p - 1) \ int p dvd R' ^ p" by (simp add: prime_dvd_mult_iff) moreover have "\int p dvd \ q" proof assume "int p dvd \ q" hence "int p \ \\ q\" using \_nz[of q] dvd_imp_le_int[of "\ q" "int p"] q by auto with p_ineqs(1) q show False by auto qed moreover have "\int p dvd fact (p - 1)" proof - have "\p dvd fact (p - 1)" using \p > 1\ p by (subst prime_dvd_fact_iff) auto hence "\int p dvd int (fact (p - 1))" by (subst int_dvd_int_iff) thus ?thesis unfolding of_nat_fact . qed moreover have "\int p dvd R' ^ p" proof assume "int p dvd R' ^ p" hence "int p dvd R'" using \prime (int p)\ prime_dvd_power by metis hence "int p \ \R'\" using \_nz[of q] dvd_imp_le_int[of R' "int p"] q by auto hence "real p \ real_of_int \R'\" by linarith also have "\ = norm (\\\Roots'. r \)" unfolding R' by simp finally show False unfolding r_def using p_ineqs(2) by linarith qed ultimately show False by blast qed have fact_p_dvd: "fact p alg_dvd h \' j" if "\' \ Roots'" "\' \ \ \ j \ p - 1" for \' j proof (cases "j \ p") case False with that have j: "j < (if \' = \ then p - 1 else p)" by auto have "h \' j = 0" unfolding h_def f_poly_altdef[OF \] by (intro poly_higher_pderiv_aux1'[OF j] dvd_smult dvd_prodI that) auto thus ?thesis by simp next case True define e where "e = (\x. if x = \ then p - 1 else p)" define Q where "Q = (\x\Roots'. [:-x, 1:] ^ e x)" define Q' where "Q' = Polynomial.smult (of_int (l^(n*p+j))) (pcompose Q [:0, 1 / of_int l:])" have "poly ((pderiv ^^ j) Q) \' / l ^ j = poly ((pderiv ^^ j) (pcompose Q [:0, 1 / of_int l:])) (l * \')" using \l > 0\ by (simp add: higher_pderiv_pcompose_linear poly_pcompose field_simps) have "sum e Roots' = (n - 1) * p + (p - 1)" unfolding e_def using \ by (subst sum.If_eq) (auto simp: card_Diff_subset n_altdef algebra_simps) also have "\ = n * p - 1" using \n > 0\ \p > 1\ by (cases n) auto finally have [simp]: "sum e Roots' = n * p - 1" . have "h \' j = of_int (l^(n*p)) * poly ((pderiv ^^ j) Q) \'" unfolding h_def f_poly_altdef[OF \] higher_pderiv_smult poly_smult e_def Q_def .. also have "poly ((pderiv ^^ j) Q) \' = of_int l ^ j * poly ((pderiv ^^ j) (pcompose Q [:0, 1 / of_int l:])) (l * \')" using \l > 0\ by (simp add: higher_pderiv_pcompose_linear poly_pcompose field_simps) also have "of_int (l ^ (n * p)) * \ = poly ((pderiv ^^ j) Q') (l * \')" by (simp add: Q'_def higher_pderiv_smult power_add) also have "fact p alg_dvd \" proof (rule fact_alg_dvd_poly_higher_pderiv) show "j \ p" by fact show "algebraic_int (of_int l * \')" by (rule alg_int) fact interpret alg_int: ring_closed "{x::complex. algebraic_int x}" by standard auto show "algebraic_int (poly.coeff Q' i)" for i proof (cases "i \ Polynomial.degree Q'") case False thus ?thesis by (simp add: coeff_eq_0) next case True hence "i \ n * p - 1" using \l > 0\ by (simp add: Q'_def degree_prod_eq Q_def degree_power_eq) also have "n * p > 0" using \n > 0\ \p > 1\ by auto hence "n * p - 1 < n * p" by simp finally have i: "i < n * p" . have "poly.coeff Q' i = of_int l ^ (n * p + j) / of_int l ^ i * poly.coeff Q i" by (simp add: Q'_def coeff_pcompose_linear field_simps) also have "of_int l ^ (n * p + j) = (of_int l ^ (n * p + j - i) :: complex) * of_int l ^ i" unfolding power_add [symmetric] using i by simp hence "of_int l ^ (n * p + j) / of_int l ^ i = (of_int l ^ (n * p + j - i) :: complex)" using \l > 0\ by (simp add: field_simps) also have "\ * poly.coeff Q i = (\X\{X. X \ (SIGMA x:Roots'. {.. i = n * p - Suc (card X)}. of_int l ^ (n * p + j - (n * p - Suc (card X))) * ((- 1) ^ card X * prod fst X))" unfolding Q_def by (subst coeff_prod_linear_factors) (auto simp: sum_distrib_left) also have "algebraic_int \" proof (intro algebraic_int_sum, goal_cases) case (1 X) hence X: "X \ (SIGMA x:Roots'. {.. card (SIGMA x:Roots'. {.. n * p - 1" using card_eq by auto also have "\ < n * p" using \n * p > 0\ by simp finally have card_less: "card X < n * p" . have "algebraic_int ((-1) ^ card X * of_int l ^ (j + 1) * (\x\X. of_int l * fst x))" using X by (intro algebraic_int_times algebraic_int_prod alg_int) auto thus ?case using card_less by (simp add: power_add prod.distrib mult_ac) qed finally show ?thesis . qed qed finally show ?thesis . qed have p_dvd: "of_nat p alg_dvd h \' j" if "\' \ Roots'" "\' \ \ \ j \ p - 1" for \' j proof - have "of_nat p alg_dvd (of_nat (fact p) :: complex)" by (intro alg_dvd_of_nat dvd_fact) (use \p > 1\ in auto) hence "of_nat p alg_dvd (fact p :: complex)" by simp also have "\ alg_dvd h \' j" using that by (intro fact_p_dvd) finally show ?thesis . qed show "fact (p - 1) alg_dvd J \" unfolding J_eq' proof (intro alg_dvd_uminus_right alg_dvd_sum, safe intro!: alg_dvd_mult algebraic_int_of_int) fix q \' j assume "q \ P" "\' \ Roots q" "j < n * p" hence "\' \ Roots'" by (auto simp: Roots'_def) show "fact (p - 1) alg_dvd h \' j" proof (cases "\' = \ \ j = p - 1") case True thus ?thesis using \fact (p - 1) alg_dvd h \ (p - 1)\ by simp next case False have "of_int (fact (p - 1)) alg_dvd (of_int (fact p) :: complex)" by (intro alg_dvd_of_int fact_dvd) auto hence "fact (p - 1) alg_dvd (fact p :: complex)" by simp also have "\ alg_dvd h \' j" using False \\' \ Roots'\ by (intro fact_p_dvd) auto finally show ?thesis . qed qed show "\of_nat p alg_dvd J \" unfolding J_eq' alg_dvd_uminus_right_iff proof (rule not_alg_dvd_sum) have "p - 1 < 1 * p" using \p > 1\ by simp also have "1 * p \ n * p" using \n > 0\ by (intro mult_right_mono) auto finally show "((q, \), p - 1) \ Sigma P Roots \ {..n > 0\ by auto next fix z assume z: "z \ Sigma P Roots \ {..),p-1)}" from z have "snd (fst z) \ Roots'" by (auto simp: Roots'_def) moreover have "fst (fst z) = q" if "\ \ Roots (fst (fst z))" proof - have "\ \ Roots (fst (fst z)) \ Roots q" "q \ P" "fst (fst z) \ P" using that q z by auto with disjoint show ?thesis unfolding disjoint_family_on_def by blast qed ultimately have "of_nat p alg_dvd h (snd (fst z)) (snd z)" using z by (intro p_dvd) auto thus "of_nat p alg_dvd (case z of (x, xa) \ (case x of (q, \') \ \i. of_int (\ q) * h \' i) xa)" using z by auto qed (use \\of_nat p alg_dvd of_int (\ q) * h \ (p-1)\ in auto) qed text \ Our next goal is to show that $J$ is rational. This is done by repeated applications of the fundamental theorem of symmetric polynomials, exploiting the fact that $J$ is symmetric in all the $\alpha_i$ for each set of conjugates. \ define g :: "int poly poly" where "g = synthetic_div (map_poly (\x. [:x:]) ((Polynomial.smult lc_factor (\P)) ^ p)) [:0, 1:]" have g: "map_poly (\p. ipoly p \) g = f_poly \" if \: "\ \ Roots'" for \ proof - interpret \: comm_ring_hom "\p. ipoly p \" by standard (auto simp: of_int_hom.poly_map_poly_eval_poly of_int_poly_hom.hom_mult) define Q :: "int poly" where "Q = (Polynomial.smult lc_factor (\P)) ^ p" have "f_poly \ = Polynomial.smult (of_int (l^(n*p))) ((\\'\Roots'. [:-\',1:])^p) div [:-\,1:]" unfolding f_poly_def div_smult_left [symmetric] prod_power_distrib[symmetric] .. also have "of_int (l^(n*p)) = (of_int l^n)^p" by (simp add: power_mult) also have "Polynomial.smult \ ((\\'\Roots'. [:-\',1:])^p) = (Polynomial.smult (of_int l ^ n) (\\'\Roots'. [:-\',1:])) ^ p" by (simp only: smult_power) also have "\ = of_int_poly Q" by (subst lc_factor) (simp_all add: Q_def of_int_poly_hom.hom_power) also have "\ div [:-\, 1:] = synthetic_div (of_int_poly Q) \" unfolding synthetic_div_altdef .. also have "\ = synthetic_div (map_poly (\p. ipoly p \) (map_poly (\x. [:x:]) Q)) (ipoly [:0, 1:] \)" by (simp add: map_poly_map_poly o_def) also have "\ = map_poly (\p. ipoly p \) g" unfolding g_def Q_def by (rule \.synthetic_div_hom) finally show ?thesis .. qed obtain Q where Q: "J \ = -(\q\P. of_int (\ q) * eval_poly of_rat (Q q) \)" if "\ \ Roots'" for \ proof - define g' :: "nat \ complex poly poly" where "g' = (\j. (map_poly of_int_poly ((pderiv ^^ j) g)))" obtain root :: "int poly \ nat \ complex" where root: "\q. q \ P \ bij_betw (root q) {..<\q} (Roots q)" using ex_bij_betw_nat_finite[OF finite_Roots] unfolding n_roots_def atLeast0LessThan by metis have "\Q'. map_poly of_rat Q' = (\x\Roots q. poly (g' j) [:x:])" if q: "q \ P" for q j proof - define Q :: "nat \ complex poly mpoly" where "Q = (\j. (\i<\q. mpoly_of_poly i (g' j)))" define ratpolys :: "complex poly set" where "ratpolys = {p. \i. poly.coeff p i \ \}" have "insertion ((\x. [:x:]) \ root q) (Q j) \ ratpolys" proof (rule symmetric_poly_of_roots_in_subring) show "ring_closed ratpolys" by standard (auto simp: ratpolys_def intro!: coeff_mult_semiring_closed) show "\m. MPoly_Type.coeff (Q j) m \ ratpolys" by (auto simp: Q_def ratpolys_def Polynomial.coeff_sum coeff_mpoly_of_poly when_def g'_def intro!: sum_in_Rats) show "vars (Q j) \ {..<\q}" unfolding Q_def by (intro order.trans[OF vars_sum] UN_least order.trans[OF vars_mpoly_of_poly]) auto show "symmetric_mpoly {..<\q} (Q j)" unfolding Q_def by (rule symmetric_mpoly_symmetric_sum[of _ id]) (auto simp: permutes_bij) interpret coeff_lift_hom: map_poly_idom_hom "\x. [:x:]" by standard define lc :: complex where "lc = of_int (Polynomial.lead_coeff q)" have "of_int_poly q = Polynomial.smult (Polynomial.lead_coeff q) (\x\Roots q. [:-x, 1:])" by (rule of_int_poly_P) fact also have "poly_lift \ = Polynomial.smult [:lc:] (\a\Roots q. [:-[:a:], 1:])" by (simp add: poly_lift_def map_poly_smult coeff_lift_hom.hom_prod lc_def) also have "(\a\Roots q. [:-[:a:], 1:]) = (\i<\q. [:-[:root q i:], 1:])" by (intro prod.reindex_bij_betw [symmetric] root q) also have "\ = (\i<\q. [:- ((\x. [:x:]) \ root q) i, 1:])" by simp finally show "poly_lift (Ring_Hom_Poly.of_int_poly q) = Polynomial.smult [:lc:] \" . have "lc \ 0" using q by (auto simp: lc_def) thus "[:inverse lc:] * [:lc:] = 1" by (simp add: field_simps) qed (auto simp: ratpolys_def coeff_pCons split: nat.splits) also have "insertion ((\x. [:x:]) \ root q) (Q j) = (\i<\q. poly (g' j) [:root q i:])" by (simp add: Q_def insertion_sum poly_sum) also have "\ = (\x\Roots q. poly (g' j) [:x:])" by (intro sum.reindex_bij_betw root q) finally have "\i. poly.coeff (\x\Roots q. poly (g' j) [:x:]) i \ \" by (auto simp: ratpolys_def) thus ?thesis using ratpolyE by metis qed then obtain Q where Q: "\q j. q \ P \ map_poly of_rat (Q q j) = (\x\Roots q. poly (g' j) [:x:])" by metis define Q' where "Q' = (\q. \j = - (\q\P. of_int (\ q) * eval_poly of_rat (Q' q) \)" if \: "\ \ Roots'" for \ proof - have "J \ = -(\q\P. of_int (\ q) * (\x\Roots q. \j)) x))" (is "_ = -?S") unfolding J_eq[OF \] .. also have "?S = (\q\P. of_int (\ q) * eval_poly of_rat (Q' q) \)" proof (rule sum.cong, goal_cases) case q: (2 q) interpret \: idom_hom "\p. ipoly p \" by standard (auto simp: of_int_hom.poly_map_poly_eval_poly of_int_poly_hom.hom_mult) have "(\x\Roots q. \j)) x) = (\jx\Roots q. poly ((pderiv ^^ j) (f_poly \)) x)" by (rule sum.swap) also have "\ = (\j)" proof (rule sum.cong, goal_cases) case j: (2 j) have "(\x\Roots q. poly ((pderiv ^^ j) (f_poly \)) x) = (\x\Roots q. poly (poly (g' j) [:x:]) \)" proof (rule sum.cong, goal_cases) case (2 x) have "poly ((pderiv ^^ j) (f_poly \)) x = poly ((pderiv ^^ j) (map_poly (\p. ipoly p \) g)) x" by (subst g[OF \, symmetric]) (rule refl) also have "\ = poly (eval_poly ((\p. [:poly p \:]) \ of_int_poly) ((pderiv ^^ j) g) [:0, 1:]) x" unfolding o_def \.map_poly_higher_pderiv [symmetric] by (simp only: \.map_poly_eval_poly) also have "\ = poly (eval_poly (\p. [:poly p \:]) (map_poly of_int_poly ((pderiv ^^ j) g)) [:0, 1:]) x" unfolding eval_poly_def by (subst map_poly_map_poly) auto also have "\ = poly (poly (map_poly of_int_poly ((pderiv ^^ j) g)) [:x:]) \" by (rule poly_poly_eq [symmetric]) also have "\ = poly (poly (g' j) [:x:]) \" by (simp add: g'_def) finally show ?case . qed auto also have "\ = poly (\x\Roots q. poly (g' j) [:x:]) \" by (simp add: poly_sum) also have "\ = eval_poly of_rat (Q q j) \" using q by (simp add: Q eval_poly_def) finally show ?case . qed auto also have "\ = eval_poly of_rat (Q' q) \" by (simp add: Q'_def of_rat_hom.eval_poly_sum) finally show ?case by simp qed auto finally show "J \ = - (\q\P. of_int (\ q) * eval_poly of_rat (Q' q) \)" . qed thus ?thesis using that[of Q'] by metis qed have "J' \ \" proof - have "(\\\Roots q. J \) \ \" if q: "q \ P" for q proof - obtain root where root: "bij_betw root {..<\q} (Roots q)" using ex_bij_betw_nat_finite[OF finite_Roots[OF q]] unfolding atLeast0LessThan n_roots_def by metis define Q' :: "complex poly" where "Q' = -(\q\P. Polynomial.smult (of_int (\ q)) (map_poly of_rat (Q q)))" have "(\\\Roots q. J \) = (\\\Roots q. -(\q\P. of_int (\ q) * eval_poly of_rat (Q q) \))" by (intro prod.cong refl Q) (auto simp: Roots'_def q) also have "\ = (\\\Roots q. poly Q' \)" by (simp add: Q'_def poly_sum eval_poly_def) also have "\ = (\i<\q. poly Q' (root i))" by (intro prod.reindex_bij_betw [symmetric] root) also have "\ = insertion root (\i<\q. mpoly_of_poly i Q')" by (simp add: insertion_prod) also have "\ \ \" proof (rule symmetric_poly_of_roots_in_subring) show "ring_closed (\ :: complex set)" by standard auto then interpret Q: ring_closed "\ :: complex set" . show "\m. MPoly_Type.coeff (\i<\q. mpoly_of_poly i Q') m \ \" by (auto intro!: Q.coeff_prod_closed sum_in_Rats simp: coeff_mpoly_of_poly when_def Q'_def Polynomial.coeff_sum) show "symmetric_mpoly {..<\q} (\i<\q. mpoly_of_poly i Q')" by (intro symmetric_mpoly_symmetric_prod'[of _ id]) (auto simp: permutes_bij) show "vars (\i<\q. mpoly_of_poly i Q') \ {..<\q}" by (intro order.trans[OF vars_prod] order.trans[OF vars_mpoly_of_poly] UN_least) auto define lc where "lc = (of_int (Polynomial.lead_coeff q) :: complex)" have "of_int_poly q = Polynomial.smult lc (\x\Roots q. [:- x, 1:])" unfolding lc_def by (rule of_int_poly_P) fact also have "(\x\Roots q. [:- x, 1:]) = (\i<\q. [:- root i, 1:])" by (intro prod.reindex_bij_betw [symmetric] root) finally show "of_int_poly q = Polynomial.smult lc (\i<\q. [:- root i, 1:])" . have "lc \ 0" using q by (auto simp: lc_def) thus "inverse lc * lc = 1" "inverse lc \ \" by (auto simp: lc_def) qed auto finally show ?thesis . qed hence "(\q\P. \\\Roots q. J \) \ \" by (rule prod_in_Rats) also have "(\q\P. \\\Roots q. J \) = J'" unfolding Roots'_def J'_def using disjoint by (intro prod.UNION_disjoint [symmetric]) (auto simp: disjoint_family_on_def) finally show "J' \ \" . qed text \ Since \J'\ is clearly an algebraic integer, we now know that it is in fact an integer. \ moreover have "algebraic_int J'" unfolding J'_def proof (intro algebraic_int_prod) fix x assume "x \ Roots'" hence "fact (p - 1) alg_dvd J x" by (intro J) thus "algebraic_int (J x)" by (rule alg_dvd_imp_algebraic_int) auto qed ultimately have "J' \ \" using rational_algebraic_int_is_int by blast text \ It is also non-zero, as none of the $J_i$ have $p$ as a factor and such cannot be zero. \ have "J' \ 0" unfolding J'_def proof (intro prod_nonzeroI) fix \ assume "\ \ Roots'" hence "\of_nat p alg_dvd J \" using J(2)[of \] by auto thus "J \ \ 0" by auto qed text \ It then clearly follows that $(p-1)!^n \leq J$: \ have "fact (p - 1) ^ n alg_dvd J'" proof - have "fact (p - 1) ^ n = (\\\Roots'. fact (p - 1))" by (simp add: n_altdef) also have "\ alg_dvd J'" unfolding J'_def by (intro prod_alg_dvd_prod J(1)) finally show ?thesis . qed have "fact (p - 1) ^ n \ norm J'" proof - from \J' \ \\ obtain J'' where [simp]: "J' = of_int J''" by (elim Ints_cases) have "of_int (fact (p - 1) ^ n) = (fact (p - 1) ^ n :: complex)" by simp also have "\ alg_dvd J'" by fact also have "J' = of_int J''" by fact finally have "fact (p - 1) ^ n dvd J''" by (subst (asm) alg_dvd_of_int_iff) moreover from \J' \ 0\ have "J'' \ 0" by auto ultimately have "\J''\ \ \fact (p - 1) ^ n\" by (intro dvd_imp_le_int) hence "real_of_int \J''\ \ real_of_int \fact (p - 1) ^ n\" by linarith also have "real_of_int \J''\ = norm J'" by simp finally show ?thesis by simp qed text \The standard M-L bound for $I_i(x)$ shows the following inequality:\ also have "norm J' \ C' * C p ^ n" proof - have "norm J' = (\x\Roots'. norm (J x))" unfolding J'_def prod_norm [symmetric] .. also have "\ \ (\x\Roots'. \q\P. real_of_int \\ q\ * (\\\Roots q. cmod \ * exp (cmod \) * C p))" proof (intro prod_mono conjI) fix x assume x: "x \ Roots'" show "norm (J x) \ (\q\P. real_of_int \\ q\ * (\\\Roots q. norm \ * exp (norm \) * C p))" unfolding J_def proof (intro sum_norm_le) fix q assume "q \ P" show "norm (of_int (\ q) * sum (I x) (Roots q)) \ real_of_int \\ q\ * (\\\Roots q. norm \ * exp (norm \) * C p)" unfolding norm_mult norm_of_int of_int_abs proof (intro mult_left_mono sum_norm_le) fix \ assume "\ \ Roots q" hence \: "\ \ Roots'" using \q \ P\ by (auto simp: Roots'_def) show "norm (I x \) \ norm \ * exp (norm \) * C p" unfolding I_def proof (intro lindemann_weierstrass_aux.lindemann_weierstrass_integral_bound) fix t assume "t \ closed_segment 0 \" also have "closed_segment 0 \ \ cball 0 R" using \R \ 0\ R_ge[OF \] by (intro closed_segment_subset) auto finally have "norm t \ R" by simp have norm_diff_le: "norm (t - y) \ 2 * R" if "y \ Roots'" for y proof - have "norm (t - y) \ norm t + norm y" by (meson norm_triangle_ineq4) also have "\ \ R + R" by (intro add_mono[OF \norm t \ R\ R_ge] that) finally show ?thesis by simp qed have "norm (poly (f_poly x) t) = \real_of_int l\ ^ (n * p) * (\y\Roots'. cmod (t - y) ^ (if y = x then p - 1 else p))" by (simp add: eval_f x f_def norm_mult norm_power flip: prod_norm) also have "\ \ \real_of_int l\ ^ (n * p) * (\y\Roots'. (2*R) ^ (if y = x then p - 1 else p))" by (intro mult_left_mono prod_mono conjI power_mono norm_diff_le) auto also have "\ = \real_of_int l\^(n*p) * (2^(p-1) * R^(p-1) * (2^p*R^p)^(n-1))" using x by (subst prod.If_eq) (auto simp: card_Diff_subset n_altdef) also have "2^(p-1) * R^(p-1) * (2^p*R^p)^(n-1) = (2^((p-1)+p*(n-1))) * (R^((p-1)+p*(n-1)))" unfolding power_mult power_mult_distrib power_add by (simp add: mult_ac) also have "(p-1)+p*(n-1) = p*n - 1" using \n > 0\ \p > 1\ by (cases n) (auto simp: algebra_simps) also have "2 ^ (p * n - 1) * R ^ (p * n - 1) = (2*R)^(n * p-1)" unfolding power_mult_distrib by (simp add: mult_ac) finally show "norm (poly (f_poly x) t) \ C p" unfolding C_def using \l > 0\ by simp qed (use \R \ 0\ \l > 0\ in \auto simp: C_def\) qed auto qed qed auto also have "\ = C' * C p ^ n" by (simp add: C'_def power_mult_distrib n_altdef flip: sum_distrib_right mult.assoc) finally show ?thesis . qed text \And with that, we have our inequality:\ finally show "fact (p - 1) ^ n \ C' * C p ^ n" . qed text \ Some simple asymptotic estimates show that this is clearly a contradiction, since the left-hand side grows much faster than the right-hand side and there are infinitely many sufficiently large primes: \ have freq: "frequently prime sequentially" using frequently_prime_cofinite unfolding cofinite_eq_sequentially . have ev: "eventually (\p. (\q\P. int p > \\ q\) \ real p > norm (\\\Roots'. of_int (l^n) * (\\'\Roots'-{\}. (\-\')))) sequentially" by (intro eventually_ball_finite \finite P\ ballI eventually_conj filterlim_real_sequentially eventually_compose_filterlim[OF eventually_gt_at_top] filterlim_int_sequentially) have "frequently (\p. fact (p - 1) ^ n \ C' * C p ^ n) sequentially" by (rule frequently_eventually_mono[OF freq ev]) (use ineq in blast) moreover have "eventually (\p. fact (p - 1) ^ n > C' * C p ^ n) sequentially" proof (cases "R = 0") case True have "eventually (\p. p * n > 1) at_top" using \n > 0\ by (intro eventually_compose_filterlim[OF eventually_gt_at_top] mult_nat_right_at_top) thus ?thesis by eventually_elim (use \n > 0\ True in \auto simp: C_def power_0_left mult_ac\) next case False hence "R > 0" using \R \ 0\ by auto define D :: real where "D = (2 * R * \real_of_int l\) ^ n" have "D > 0" using \R > 0\ \l > 0\ unfolding D_def by (intro zero_less_power) auto have "(\p. C' * C p ^ n) \ O(\p. C p ^ n)" by simp also have "(\p. C p ^ n) \ O(\p. ((2 * R * l) ^ (n * p)) ^ n)" proof (rule landau_o.big_power[OF bigthetaD1]) have np: "eventually (\p. p * n > 0) at_top" using \n > 0\ by (intro eventually_compose_filterlim[OF eventually_gt_at_top] mult_nat_right_at_top) have "eventually (\p. (2 * R) * C p = (2 * R * l) ^ (n * p)) at_top" using np proof eventually_elim case (elim p) have "2 * R * C p = l ^ (n * p) * (2 * R) ^ (Suc (n * p - 1))" by (simp add: C_def algebra_simps) also have "Suc (n * p - 1) = n * p" using elim by auto finally show ?case by (simp add: algebra_simps) qed hence "(\p. (2 * R) * C p) \ \(\p. (2 * R * l) ^ (n * p))" by (intro bigthetaI_cong) thus "C \ \(\p. (2 * R * l) ^ (n * p))" using \R > 0\ by simp qed also have "\ = O(\p. (D ^ p) ^ n)" using \l > 0\ by (simp flip: power_mult add: power2_eq_square mult_ac D_def) also have "(\p. (D ^ p) ^ n) \ o(\p. fact (p - 1) ^ n)" proof (intro landau_o.small_power) have "eventually (\p. D ^ p = D * D ^ (p - 1)) at_top" using eventually_gt_at_top[of 0] by eventually_elim (use \D > 0\ in \auto simp flip: power_Suc\) hence "(\p. D ^ p) \ \(\p. D * D ^ (p - 1))" by (intro bigthetaI_cong) hence "(\p. D ^ p) \ \(\p. D ^ (p - 1))" using \D > 0\ by simp also have "(\p. D ^ (p - 1)) \ o(\p. fact (p - 1))" by (intro smalloI_tendsto[OF filterlim_compose[OF power_over_fact_tendsto_0]] - filterlim_minus_nat_at_top) auto + filterlim_minus_const_nat_at_top) auto finally show "(\p. D ^ p) \ o(\x. fact (x - 1))" . qed fact+ finally have smallo: "(\p. C' * C p ^ n) \ o(\p. fact (p - 1) ^ n)" . have "eventually (\p. \C' * C p ^ n\ \ 1/2 * fact (p - 1) ^ n) at_top" using landau_o.smallD[OF smallo, of "1/2"] by simp thus "eventually (\p. C' * C p ^ n < fact (p - 1) ^ n) at_top" proof eventually_elim case (elim p) have "C' * C p ^ n \ \C' * C p ^ n\" by simp also have "\ \ 1/2 * fact (p - 1) ^ n" by fact also have "\ < fact (p - 1) ^ n" by simp finally show ?case . qed qed ultimately have "frequently (\p::nat. False) sequentially" by (rule frequently_eventually_mono) auto thus False by simp qed subsection \Removing the restriction of full sets of conjugates\ text \ We will now remove the restriction that the $\alpha_i$ must occur in full sets of conjugates by multiplying the equality with all permutations of roots. \ lemma Hermite_Lindemann_aux2: fixes X :: "complex set" and \ :: "complex \ int" assumes "finite X" assumes nz: "\x. x \ X \ \ x \ 0" assumes alg: "\x. x \ X \ algebraic x" assumes sum0: "(\x\X. of_int (\ x) * exp x) = 0" shows "X = {}" proof (rule ccontr) assume "X \ {}" note [intro] = \finite X\ text \ Let \P\ be the smallest integer polynomial whose roots are a superset of \X\: \ define P :: "int poly" where "P = \(min_int_poly ` X)" define Roots :: "complex set" where "Roots = {x. ipoly P x = 0}" have [simp]: "P \ 0" using \finite X\ by (auto simp: P_def) have [intro]: "finite Roots" unfolding Roots_def by (intro poly_roots_finite) auto have "X \ Roots" proof safe fix x assume "x \ X" hence "ipoly (min_int_poly x) x = 0" by (intro ipoly_min_int_poly alg) thus "x \ Roots" using \finite X\ \x \ X\ by (auto simp: Roots_def P_def of_int_poly_hom.hom_prod poly_prod) qed have "squarefree (of_int_poly P :: complex poly)" unfolding P_def of_int_poly_hom.hom_prod proof (rule squarefree_prod_coprime; safe) fix x assume "x \ X" thus "squarefree (of_int_poly (min_int_poly x) :: complex poly)" by (intro squarefree_of_int_polyI) auto next fix x y assume xy: "x \ X" "y \ X" "min_int_poly x \ min_int_poly y" thus "Rings.coprime (of_int_poly (min_int_poly x)) (of_int_poly (min_int_poly y) :: complex poly)" by (intro coprime_of_int_polyI[OF primes_coprime]) auto qed text \ Since we will need a numbering of these roots, we obtain one: \ define n where "n = card Roots" obtain Root where Root: "bij_betw Root {..finite Roots\] unfolding n_def atLeast0LessThan by metis define unRoot :: "complex \ nat" where "unRoot = inv_into {.. Roots" for x unfolding unRoot_def using Root that by (subst f_inv_into_f) (auto simp: bij_betw_def) have [simp, intro]: "Root i \ Roots" if "i < n" for i using Root that by (auto simp: bij_betw_def) have [simp, intro]: "unRoot x < n" if "x \ Roots" for x using unRoot that by (auto simp: bij_betw_def) text \ We will also need to convert between permutations of natural numbers less than \n\ and permutations of the roots: \ define convert_perm :: "(nat \ nat) \ (complex \ complex)" where "convert_perm = (\\ x. if x \ Roots then Root (\ (unRoot x)) else x)" have bij_convert: "bij_betw convert_perm {\. \ permutes {... \ permutes Roots}" using bij_betw_permutations[OF Root] unfolding convert_perm_def unRoot_def . have permutes_convert_perm [intro]: "convert_perm \ permutes Roots" if "\ permutes {.. using that bij_convert unfolding bij_betw_def by blast have convert_perm_compose: "convert_perm (\ \ \) = convert_perm \ \ convert_perm \" if "\ permutes {.. permutes {.. \ proof (intro ext) fix x show "convert_perm (\ \ \) x = (convert_perm \ \ convert_perm \) x" proof (cases "x \ Roots") case True thus ?thesis using permutes_in_image[OF that(2), of "unRoot x"] by (auto simp: convert_perm_def bij_betw_def) qed (auto simp: convert_perm_def) qed text \ We extend the coefficient vector to the new roots by setting their coefficients to 0: \ define \' where "\' = (\x. if x \ X then \ x else 0)" text \ We now define the set of all permutations of our roots: \ define perms where "perms = {\. \ permutes Roots}" have [intro]: "finite perms" unfolding perms_def by (rule finite_permutations) auto have [simp]: "card perms = fact n" unfolding perms_def n_def by (intro card_permutations) auto text \ The following is the set of all \n!\-tuples of roots, disregarding permutation of components. In other words: all multisets of roots with size \n!\. \ define Roots_ms :: "complex multiset set" where "Roots_ms = {X. set_mset X \ Roots \ size X = fact n}" have [intro]: "finite Roots_ms" unfolding Roots_ms_def by (rule finite_multisets_of_size) auto text \ Next, the following is the set of \n!\-tuples whose entries are precisely the multiset \X\: \ define tuples :: "complex multiset \ ((complex \ complex) \ complex) set" where "tuples = (\X. {f\perms \\<^sub>E Roots. image_mset f (mset_set perms) = X})" have fin_tuples [intro]: "finite (tuples X)" for X unfolding tuples_def by (rule finite_subset[of _ "perms \\<^sub>E Roots", OF _ finite_PiE]) auto define tuples' :: "(complex multiset \ ((complex \ complex) \ complex)) set" where "tuples' = (SIGMA X:Roots_ms. tuples X)" text \ The following shows that our \<^term>\tuples\ definition is stable under permutation of the roots. \ have bij_convert': "bij_betw (\f. f \ (\g. \ \ g)) (tuples X) (tuples X)" if \: "\ permutes Roots" for \ X proof (rule bij_betwI) have *: "(\f. f \ (\) \) \ tuples X \ tuples X" if \: "\ permutes Roots" for \ proof fix f assume f: "f \ tuples X" show "f \ (\) \ \ tuples X" unfolding tuples_def proof safe fix \' assume \': "\' \ perms" show "(f \ (\) \) \' \ Roots" using permutes_compose[OF _ \, of \'] \ \' f by (auto simp: perms_def tuples_def) next fix \' assume \': "\' \ perms" have "\(\ \ \') permutes Roots" proof assume "(\ \ \') permutes Roots" hence "inv_into UNIV \ \ (\ \ \') permutes Roots" by (rule permutes_compose) (use permutes_inv[OF \] in simp_all) also have "inv_into UNIV \ \ (\ \ \') = \'" by (auto simp: fun_eq_iff permutes_inverses[OF \]) finally show False using \' by (simp add: perms_def) qed thus "(f \ (\) \) \' = undefined" using f by (auto simp: perms_def tuples_def) next have "image_mset (f \ (\) \) (mset_set perms) = image_mset f (image_mset ((\) \) (mset_set perms))" by (rule multiset.map_comp [symmetric]) also have "image_mset ((\) \) (mset_set perms) = mset_set perms" using bij_betw_permutes_compose_left[OF \] by (subst image_mset_mset_set) (auto simp: bij_betw_def perms_def) also have "image_mset f \ = X" using f by (auto simp: tuples_def) finally show "image_mset (f \ (\) \) (mset_set perms) = X" . qed qed show "(\f. f \ (\) \) \ tuples X \ tuples X" by (rule *) fact show "(\f. f \ (\) (inv_into UNIV \)) \ tuples X \ tuples X" by (intro * permutes_inv) fact show "f \ (\) \ \ (\) (inv_into UNIV \) = f" if "f \ tuples X" for f by (auto simp: fun_eq_iff o_def permutes_inverses[OF \]) show "f \ (\) (inv_into UNIV \) \ (\) \ = f" if "f \ tuples X" for f by (auto simp: fun_eq_iff o_def permutes_inverses[OF \]) qed text \ Next, we define the multiset of of possible exponents that we can get for a given \n!\-multiset of roots, \ define R :: "complex multiset \ complex multiset" where "R = (\X. image_mset (\f. \\\perms. \ (f \)) (mset_set (tuples X)))" text \ We show that, for each such multiset, there is a content-free integer polynomial that has exactly these exponents as roots. This shows that they form a full set of conjugates (but note this polynomial is not necessarily squarefree). The proof is yet another application of the fundamental theorem of symmetric polynomials. \ obtain Q :: "complex multiset \ int poly" where Q: "\X. X \ Roots_ms \ poly_roots (of_int_poly (Q X)) = R X" "\X. X \ Roots_ms \ content (Q X) = 1" proof - { fix X :: "complex multiset" assume X: "X \ Roots_ms" define Q :: "complex poly mpoly" where "Q = (\f\tuples X. Const [:0, 1:] - (\\ | \ permutes {.. (unRoot (f (convert_perm \))))))" define Q1 where "Q1 = (\f\tuples X. [:- (\\ | \ permutes Roots. \ (f \)), 1:])" define ratpolys :: "complex poly set" where "ratpolys = {p. \i. poly.coeff p i \ \}" have "insertion (\x. [:Root x:]) Q \ ratpolys" proof (rule symmetric_poly_of_roots_in_subring[where l = "\x. [:x:]"]) show "ring_closed ratpolys" unfolding ratpolys_def by standard (auto intro: coeff_mult_semiring_closed) then interpret ratpolys: ring_closed ratpolys . have "pCons 0 1 \ ratpolys " by (auto simp: ratpolys_def coeff_pCons split: nat.splits) thus "\m. MPoly_Type.coeff Q m \ ratpolys" unfolding Q_def by (intro allI ratpolys.coeff_prod_closed) (auto intro!: ratpolys.minus_closed ratpolys.sum_closed ratpolys.uminus_closed simp: coeff_Var mpoly_coeff_Const when_def) next show "ring_homomorphism (\x::complex. [:x:])" .. next have "\ (unRoot (f (convert_perm \))) < n" if "f \ tuples X" "\ permutes {.. proof - have "convert_perm \ \ perms" using bij_convert that(2) by (auto simp: bij_betw_def perms_def) hence "f (convert_perm \) \ Roots" using that by (auto simp: tuples_def) thus ?thesis using permutes_in_image[OF that(2)] by simp qed thus "vars Q \ {.. ratpolys" by (auto simp: ratpolys_def coeff_pCons lc_def split: nat.splits) show "\i. [:poly.coeff (of_int_poly P) i:] \ ratpolys" by (auto simp: ratpolys_def coeff_pCons split: nat.splits) have "lc \ 0" by (auto simp: lc_def) thus "[:inverse lc:] * [:lc:] = 1" by auto have "rsquarefree (of_int_poly P :: complex poly)" using \squarefree (of_int_poly P :: complex poly)\ by (intro squarefree_imp_rsquarefree) hence "of_int_poly P = Polynomial.smult lc (\x\Roots. [:-x, 1:])" unfolding lc_def Roots_def of_int_hom.hom_lead_coeff[symmetric] by (rule complex_poly_decompose_rsquarefree [symmetric]) also have "(\x\Roots. [:-x, 1:]) = (\ii assume \: "\ permutes {.. Q = (\f\tuples X. Const (pCons 0 1) - (\ \ | \ permutes {.. \ \) (unRoot (f (convert_perm \))))))" by (simp add: Q_def permutes_bij[OF \]) also have "\ = (\f\tuples X. Const (pCons 0 1) - (\ \ | \ permutes {.. \ \) (unRoot ((f \ (\\. convert_perm \ \ \)) (convert_perm \))))))" using \ by (intro prod.reindex_bij_betw [OF bij_convert', symmetric]) auto also have "\ = Q" unfolding Q_def proof (rule prod.cong, goal_cases) case (2 f) have "(\ \ | \ permutes {.. \ \) (unRoot ((f \ (\\. convert_perm \ \ \)) (convert_perm \))))) = (\ \ | \ permutes {.. \ \) (unRoot (f (convert_perm (\ \ \))))))" using \ by (intro sum.cong refl, subst convert_perm_compose) simp_all also have "\ = (\ \ | \ permutes {.. (unRoot (f (convert_perm \)))))" using \ by (rule setum_permutations_compose_left [symmetric]) finally show ?case by simp qed auto finally show "mpoly_map_vars \ Q = Q" . qed qed auto also have "insertion (\x. [:Root x:]) Q = Q1" unfolding Q_def Q1_def insertion_prod insertion_sum insertion_diff insertion_Const insertion_Var proof (intro prod.cong, goal_cases) case f: (2 f) have "(\\ | \ permutes {.. (unRoot (f (convert_perm \)))):]) = (\\ | \ permutes {.. (f (convert_perm \)):])" proof (rule sum.cong, goal_cases) case (2 \) have "convert_perm \ permutes Roots" using bij_convert 2 by (auto simp: bij_betw_def) hence "f (convert_perm \) \ Roots" using f by (auto simp: tuples_def perms_def) thus ?case by (simp add: convert_perm_def) qed simp_all also have "\ = (\\ | \ permutes Roots. [:\ (f \):])" by (rule sum.reindex_bij_betw[OF bij_convert]) finally show ?case by (simp flip: pCons_one coeff_lift_hom.hom_sum) qed simp_all finally have "Q1 \ ratpolys" by auto then obtain Q2 :: "rat poly" where Q2: "Q1 = map_poly of_rat Q2" unfolding ratpolys_def using ratpolyE[of Q1] by blast have "Q1 \ 0" unfolding Q1_def using fin_tuples[of X] by auto with Q2 have "Q2 \ 0" by auto obtain Q3 :: "int poly" and lc :: rat where Q3: "Q2 = Polynomial.smult lc (of_int_poly Q3)" and "lc > 0" and "content Q3 = 1" using rat_to_normalized_int_poly_exists[OF \Q2 \ 0\] by metis have "poly_roots (of_int_poly Q3) = poly_roots (map_poly (of_rat \ of_int) Q3)" by simp also have "map_poly (of_rat \ of_int) Q3 = map_poly of_rat (map_poly of_int Q3)" by (subst map_poly_map_poly) auto also have "poly_roots \ = poly_roots (Polynomial.smult (of_rat lc) \)" using \lc > 0\ by simp also have "Polynomial.smult (of_rat lc) (map_poly of_rat (map_poly of_int Q3)) = map_poly of_rat (Polynomial.smult lc (map_poly of_int Q3))" by (simp add: of_rat_hom.map_poly_hom_smult) also have "\ = Q1" by (simp only: Q3 [symmetric] Q2 [symmetric]) also have "poly_roots Q1 = R X" unfolding Q1_def by (subst poly_roots_prod, force, subst poly_roots_linear) (auto simp: R_def perms_def sum_mset_image_mset_singleton sum_unfold_sum_mset) finally have "\Q. poly_roots (of_int_poly Q) = R X \ content Q = 1" using \content Q3 = 1\ by metis } hence "\Q. \X\Roots_ms. poly_roots (of_int_poly (Q X)) = R X \ content (Q X) = 1" by metis thus ?thesis using that by metis qed text \ We can now collect all the $e^{\sum \alpha_i}$ that happen to be equal and let the following be their coefficients: \ define \'' :: "int poly \ int" where "\'' = (\q. \X\Roots_ms. int (count (prime_factorization (Q X)) q) * (\x\#X. \' x))" have supp_\'': "{q. \'' q \ 0} \ (\X\Roots_ms. prime_factors (Q X))" unfolding \''_def using sum.not_neutral_contains_not_neutral by fastforce text \ We have to prove that \\''\ is not zero everywhere. We do this by selecting the nonzero term with the maximal exponent (w.r.t. the lexicographic ordering on the complex numbers) in every factor of the product and show that there is no other summand corresponding to these, so that their non-zero coefficient cannot get cancelled. \ have "{q. \'' q \ 0} \ {}" proof - define f where "f = restrict (\\. inv_into UNIV \ (complex_lex.Max (\ ` X))) perms" have f: "f \ perms \ X" proof fix \ assume \: "\ \ perms" have "complex_lex.Max (\ ` X) \ \ ` X" using \X \ {}\ by (intro complex_lex.Max_in finite_imageI) auto thus "f \ \ X" using \ by (auto simp: f_def permutes_inverses[of \ Roots] perms_def) qed hence f': "f \ perms \\<^sub>E Roots" using \X \ Roots\ by (auto simp: f_def PiE_def) define Y where "Y = image_mset f (mset_set perms)" have "Y \ Roots_ms" using f' \finite perms\ by (auto simp: Roots_ms_def Y_def) have "(\\\perms. \ (f \)) \# R Y" proof - from f' have "f \ tuples Y" unfolding tuples_def Y_def by simp thus ?thesis unfolding R_def using fin_tuples[of Y] by auto qed also have "R Y = poly_roots (of_int_poly (Q Y))" by (rule Q(1) [symmetric]) fact also have "\ = (\p\#prime_factorization (Q Y). poly_roots (of_int_poly p))" by (rule poly_roots_of_int_conv_sum_prime_factors) finally obtain q where q: "q \ prime_factors (Q Y)" "(\\\perms. \ (f \)) \# poly_roots (of_int_poly q)" by auto have "\'' q = (\X\{Y}. int (count (prime_factorization (Q X)) q) * prod_mset (image_mset \' X))" unfolding \''_def proof (intro sum.mono_neutral_right ballI) fix Y' assume Y': "Y' \ Roots_ms - {Y}" show "int (count (prime_factorization (Q Y')) q) * \\<^sub># (image_mset \' Y') = 0" proof (cases "set_mset Y' \ X") case Y'_subset: True have "q \ prime_factors (Q Y')" proof assume q': "q \ prime_factors (Q Y')" have "poly_roots (of_int_poly q :: complex poly) \# poly_roots (of_int_poly (Q Y'))" using q' by (intro dvd_imp_poly_roots_subset of_int_poly_hom.hom_dvd) auto with q(2) have "(\\\perms. \ (f \)) \# poly_roots (of_int_poly (Q Y'))" by (meson mset_subset_eqD) also have "poly_roots (of_int_poly (Q Y')) = R Y'" using Q(1)[of Y'] Y' by auto finally obtain g where g: "g \ tuples Y'" "(\\\perms. \ (f \)) = (\\\perms. \ (g \))" unfolding R_def using fin_tuples[of Y'] by auto moreover have "(\\\perms. \ (g \)) <\<^sub>\ (\\\perms. \ (f \))" proof (rule sum_strict_mono_ex1_complex_lex) show le: "\\\perms. \ (g \) \\<^sub>\ \ (f \)" proof fix \ assume \: "\ \ perms" hence \': "\ permutes Roots" by (auto simp: perms_def) have "image_mset g (mset_set perms) = Y'" using g by (auto simp: tuples_def) also have "set_mset \ \ X" by fact finally have "g ` perms \ X" using \finite perms\ by auto hence "\ (g \) \\<^sub>\ complex_lex.Max (\ ` X)" using \finite perms\ \ by (intro complex_lex.Max.coboundedI finite_imageI imageI) (auto simp: tuples_def) also have "\ = \ (f \)" using \ by (simp add: f_def permutes_inverses[OF \']) finally show "\ (g \) \\<^sub>\ \ (f \)" . qed have "image_mset g (mset_set perms) \ image_mset f (mset_set perms)" using Y' g by (auto simp: tuples_def Y_def) then obtain \ where \: "\ \# mset_set perms" "g \ \ f \" by (meson multiset.map_cong) have "\ permutes Roots" using \ \finite perms\ by (auto simp: perms_def) have "\ (g \) \ \ (f \)" using permutes_inj[OF \\ permutes Roots\] \ by (auto simp: inj_def) moreover have "\ (g \) \\<^sub>\ \ (f \)" using le \ \finite perms\ by auto ultimately have "\ (g \) <\<^sub>\ \ (f \)" by simp thus "\\\perms. \ (g \) <\<^sub>\ \ (f \)" using \ \finite perms\ by auto qed (use \finite perms\ in simp_all) ultimately show False by simp qed thus ?thesis by auto qed (auto simp: \'_def) qed (use \Y \ Roots_ms\ in auto) also have "\ = int (count (prime_factorization (Q Y)) q) * prod_mset (image_mset \' Y)" by simp also have "\ \ 0" using q nz \finite X\ \X \ {}\ \finite perms\ f by (auto simp: \'_def Y_def) finally show "{q. \'' q \ 0} \ {}" by auto qed text \ We are now ready for the final push: we start with the original sum that we know to be zero, multiply it with the other permutations, and then multiply out the sum. \ have "0 = (\x\X. \ x * exp x)" using sum0 .. also have "\ = (\x\Roots. \' x * exp x)" by (intro sum.mono_neutral_cong_left \X \ Roots\) (auto simp: \'_def) also have "\ dvd (\\\perms. \x\Roots. \' x * exp (\ x))" by (rule dvd_prodI[OF \finite perms\]) (use permutes_id[of Roots] in \simp_all add: id_def perms_def\) also have "\ = (\f\perms \\<^sub>E Roots. \\\perms. \' (f \) * exp (\ (f \)))" by (rule prod_sum_PiE) auto also have "\ = (\f\perms \\<^sub>E Roots. (\\\perms. \' (f \)) * exp (\\\perms. \ (f \)))" using \finite perms\ by (simp add: prod.distrib exp_sum) also have "\ = (\(X,f)\tuples'. (\\\perms. \' (f \)) * exp (\\\perms. \ (f \)))" using \finite perms\ by (intro sum.reindex_bij_witness[of _ snd "\f. (image_mset f (mset_set perms), f)"]) (auto simp: tuples'_def tuples_def Roots_ms_def PiE_def Pi_def) also have "\ = (\(X,f)\tuples'. (\x\#X. \' x) * exp (\\\perms. \ (f \)))" proof (safe intro!: sum.cong) fix X :: "complex multiset" and f :: "(complex \ complex) \ complex" assume "(X, f) \ tuples'" hence X: "X \ Roots_ms" "X = image_mset f (mset_set perms)" and f: "f \ perms \\<^sub>E Roots" by (auto simp: tuples'_def tuples_def) have "(\\\perms. \' (f \)) = (\\\#mset_set perms. \' (f \))" by (meson prod_unfold_prod_mset) also have "\ = (\x\#X. \' x)" unfolding X(2) by (simp add: multiset.map_comp o_def) finally show "(\\\perms. \' (f \)) * exp (\\\perms. \ (f \)) = (\x\#X. \' x) * exp (\\\perms. \ (f \))" by simp qed also have "\ = (\X\Roots_ms. \f\tuples X. (\x\#X. \' x) * exp (\\\perms. \ (f \)))" unfolding tuples'_def by (intro sum.Sigma [symmetric]) auto also have "\ = (\X\Roots_ms. of_int (\x\#X. \' x) * (\f\tuples X. exp (\\\perms. \ (f \))))" by (simp add: sum_distrib_left) also have "\ = (\X\Roots_ms. of_int (\x\#X. \' x) * (\x\#R X. exp x))" by (simp only: R_def multiset.map_comp o_def sum_unfold_sum_mset) also have "\ = (\X\Roots_ms. of_int (\x\#X. \' x) * (\x\#poly_roots (of_int_poly (Q X)). exp x))" by (intro sum.cong) (simp_all flip: Q) text \ Our problem now is that the polynomials \Q X\ can still contain multiple roots and that their roots might not be disjoint. We therefore split them all into irreducible factors and collect equal terms. \ also have "\ = (\X\Roots_ms. (\p. of_int (int (count (prime_factorization (Q X)) p) * (\x\#X. \' x)) * (\x | ipoly p x = 0. exp x)))" proof (rule sum.cong, goal_cases) case (2 X) have "(\x\#poly_roots (of_int_poly (Q X) :: complex poly). exp x) = (\x \# (\p\#prime_factorization (Q X). poly_roots (of_int_poly p)). exp x)" by (subst poly_roots_of_int_conv_sum_prime_factors) (rule refl) also have "\ = (\p\#prime_factorization (Q X). \x\#poly_roots (of_int_poly p). exp x)" by (rule sum_mset_image_mset_sum_mset_image_mset) also have "rsquarefree (of_int_poly p :: complex poly)" if "p \ prime_factors (Q X)" for p proof (rule irreducible_imp_rsquarefree_of_int_poly) have "prime p" using that by auto thus "irreducible p" by blast next show "Polynomial.degree p > 0" by (intro content_1_imp_nonconstant_prime_factors[OF Q(2) that] 2) qed hence "(\p\#prime_factorization (Q X). \x\#poly_roots (of_int_poly p). exp x) = (\p\#prime_factorization (Q X). \x | ipoly p x = 0. exp (x :: complex))" unfolding sum_unfold_sum_mset by (intro arg_cong[of _ _ sum_mset] image_mset_cong sum.cong refl, subst rsquarefree_poly_roots_eq) auto also have "\ = (\p. count (prime_factorization (Q X)) p * (\x | ipoly p x = 0. exp (x :: complex)))" by (rule sum_mset_conv_Sum_any) also have "of_int (\x\#X. \' x) * \ = (\p. of_int (int (count (prime_factorization (Q X)) p) * (\x\#X. \' x)) * (\x | ipoly p x = 0. exp x))" by (subst Sum_any_right_distrib) (auto simp: mult_ac) finally show ?case by simp qed auto also have "\ = (\q. of_int (\'' q) * (\x | ipoly q x = 0. exp x))" unfolding \''_def of_int_sum by (subst Sum_any_sum_swap [symmetric]) (auto simp: sum_distrib_right) also have "\ = (\q | \'' q \ 0. of_int (\'' q) * (\x | ipoly q x = 0. exp x))" by (intro Sum_any.expand_superset finite_subset[OF supp_\'']) auto finally have "(\q | \'' q \ 0. of_int (\'' q) * (\x | ipoly q x = 0. exp (x :: complex))) = 0" by simp text \ We are now in the situation of our the specialised Hermite--Lindemann Theorem we proved earlier and can easily derive a contradiction. \ moreover have "(\q | \'' q \ 0. of_int (\'' q) * (\x | ipoly q x = 0. exp (x :: complex))) \ 0" proof (rule Hermite_Lindemann_aux1) show "finite {q. \'' q \ 0}" by (rule finite_subset[OF supp_\'']) auto next show "pairwise Rings.coprime {q. \'' q \ 0}" proof (rule pairwiseI, clarify) fix p q assume pq: "p \ q" "\'' p \ 0" "\'' q \ 0" hence "prime p" "prime q" using supp_\'' Q(2) by auto with pq show "Rings.coprime p q" by (simp add: primes_coprime) qed next fix q :: "int poly" assume q: "q \ {q. \'' q \ 0}" also note supp_\'' finally obtain X where X: "X \ Roots_ms" "q \ prime_factors (Q X)" by blast show "irreducible q" using X by (intro prime_elem_imp_irreducible prime_imp_prime_elem) auto show "Polynomial.degree q > 0" using X by (intro content_1_imp_nonconstant_prime_factors[OF Q(2)[of X]]) qed (use \{x. \'' x \ 0} \ {}\ in auto) ultimately show False by contradiction qed subsection \Removing the restriction to integer coefficients\ text \ Next, we weaken the restriction that the $\beta_i$ must be integers to the restriction that they must be rationals. This is done simply by multiplying with the least common multiple of the demoninators. \ lemma Hermite_Lindemann_aux3: fixes X :: "complex set" and \ :: "complex \ rat" assumes "finite X" assumes nz: "\x. x \ X \ \ x \ 0" assumes alg: "\x. x \ X \ algebraic x" assumes sum0: "(\x\X. of_rat (\ x) * exp x) = 0" shows "X = {}" proof - define l :: int where "l = Lcm ((snd \ quotient_of \ \) ` X)" have [simp]: "snd (quotient_of r) \ 0" for r using quotient_of_denom_pos'[of r] by simp have [simp]: "l \ 0" using \finite X\ by (auto simp: l_def Lcm_0_iff) have "of_int l * \ x \ \" if "x \ X" for x proof - define a b where "a = fst (quotient_of (\ x))" and "b = snd (quotient_of (\ x))" have "b > 0" using quotient_of_denom_pos'[of "\ x"] by (auto simp: b_def) have "\ x = of_int a / of_int b" by (intro quotient_of_div) (auto simp: a_def b_def) also have "of_int l * \ = of_int (l * a) / of_int b" using \b > 0\ by (simp add: field_simps) also have "\ \ \" using that by (intro of_int_divide_in_Ints) (auto simp: l_def b_def) finally show ?thesis . qed hence "\x\X. \n. of_int n = of_int l * \ x" using Ints_cases by metis then obtain \' where \': "of_int (\' x) = of_int l * \ x" if "x \ X" for x by metis show ?thesis proof (rule Hermite_Lindemann_aux2) have "0 = of_int l * (\x\X. of_rat (\ x) * exp x :: complex)" by (simp add: sum0) also have "\ = (\x\X. of_int (\' x) * exp x)" unfolding sum_distrib_left proof (rule sum.cong, goal_cases) case (2 x) have "of_int l * of_rat (\ x) = of_rat (of_int l * \ x)" by (simp add: of_rat_mult) also have "of_int l * \ x = of_int (\' x)" using 2 by (rule \' [symmetric]) finally show ?case by (simp add: mult_ac) qed simp_all finally show "\ = 0" .. next fix x assume "x \ X" hence "of_int (\' x) \ (0 :: rat)" using nz by (subst \') auto thus "\' x \ 0" by auto qed (use alg \finite X\ in auto) qed text \ Next, we weaken the restriction that the $\beta_i$ must be rational to them being algebraic. Similarly to before, this is done by multiplying over all possible permutations of the $\beta_i$ (in some sense) to introduce more symmetry, from which it then follows by the fundamental theorem of symmetric polynomials that the resulting coefficients are rational. \ lemma Hermite_Lindemann_aux4: fixes \ :: "complex \ complex" assumes [intro]: "finite X" assumes alg1: "\x. x \ X \ algebraic x" assumes alg2: "\x. x \ X \ algebraic (\ x)" assumes nz: "\x. x \ X \ \ x \ 0" assumes sum0: "(\x\X. \ x * exp x) = 0" shows "X = {}" proof (rule ccontr) assume X: "X \ {}" note [intro!] = finite_PiE text \ We now take more or less the same approach as before, except that now we find a polynomial that has all of the conjugates of the coefficients \\\ as roots. Note that this is a slight deviation from Baker's proof, who picks one polynomial for each \\\ independently. I did it this way because, as Bernard~\<^cite>\"bernard"\ observed, it makes the proof a bit easier. \ define P :: "int poly" where "P = \((min_int_poly \ \) ` X)" define Roots :: "complex set" where "Roots = {x. ipoly P x = 0}" have "0 \ Roots" using \finite X\ alg2 nz by (auto simp: Roots_def P_def poly_prod) have [simp]: "P \ 0" using \finite X\ by (auto simp: P_def) have [intro]: "finite Roots" unfolding Roots_def by (intro poly_roots_finite) auto have "\ ` X \ Roots" proof safe fix x assume "x \ X" hence "ipoly (min_int_poly (\ x)) (\ x) = 0" by (intro ipoly_min_int_poly alg2) thus "\ x \ Roots" using \finite X\ \x \ X\ by (auto simp: Roots_def P_def of_int_poly_hom.hom_prod poly_prod) qed have "squarefree (of_int_poly P :: complex poly)" unfolding P_def of_int_poly_hom.hom_prod o_def proof (rule squarefree_prod_coprime; safe) fix x assume "x \ X" thus "squarefree (of_int_poly (min_int_poly (\ x)) :: complex poly)" by (intro squarefree_of_int_polyI) auto next fix x y assume xy: "x \ X" "y \ X" "min_int_poly (\ x) \ min_int_poly (\ y)" thus "Rings.coprime (of_int_poly (min_int_poly (\ x))) (of_int_poly (min_int_poly (\ y)) :: complex poly)" by (intro coprime_of_int_polyI[OF primes_coprime]) auto qed define n where "n = card Roots" define m where "m = card X" have "Roots \ {}" using \\ ` X \ Roots\ \X \ {}\ by auto hence "n > 0" "m > 0" using \finite Roots\ \finite X\ \X \ {}\ by (auto simp: n_def m_def) have fin1 [simp]: "finite (X \\<^sub>E Roots)" by auto have [simp]: "card (X \\<^sub>E Roots) = n ^ m" by (subst card_PiE) (auto simp: m_def n_def) text \ We again find a bijection between the roots and the natural numbers less than \n\: \ obtain Root where Root: "bij_betw Root {..finite Roots\] unfolding n_def atLeast0LessThan by metis define unRoot :: "complex \ nat" where "unRoot = inv_into {.. Roots" for x unfolding unRoot_def using Root that by (subst f_inv_into_f) (auto simp: bij_betw_def) have [simp, intro]: "Root i \ Roots" if "i < n" for i using Root that by (auto simp: bij_betw_def) have [simp, intro]: "unRoot x < n" if "x \ Roots" for x using unRoot that by (auto simp: bij_betw_def) text \ And we again define the set of multisets and tuples that we will get in the expanded product. \ define Roots_ms :: "complex multiset set" where "Roots_ms = {Y. set_mset Y \ X \ size Y = n ^ m}" have [intro]: "finite Roots_ms" unfolding Roots_ms_def by (rule finite_multisets_of_size) auto define tuples :: "complex multiset \ ((complex \ complex) \ complex) set" where "tuples = (\Y. {f\(X \\<^sub>E Roots) \\<^sub>E X. image_mset f (mset_set (X \\<^sub>E Roots)) = Y})" have [intro]: "finite (tuples Y)" for Y unfolding tuples_def by (rule finite_subset[of _ "(X \\<^sub>E Roots) \\<^sub>E X"]) auto text \ We will also need to convert permutations over the natural and over the roots again. \ define convert_perm :: "(nat \ nat) \ (complex \ complex)" where "convert_perm = (\\ x. if x \ Roots then Root (\ (unRoot x)) else x)" have bij_convert: "bij_betw convert_perm {\. \ permutes {... \ permutes Roots}" using bij_betw_permutations[OF Root] unfolding convert_perm_def unRoot_def . have permutes_convert_perm [intro]: "convert_perm \ permutes Roots" if "\ permutes {.. using that bij_convert unfolding bij_betw_def by blast text \ We also need a small lemma showing that our tuples are stable under permutation of the roots. \ have bij_betw_compose_perm: "bij_betw (\f. restrict (\g. f (restrict (\ \ g) X)) (X \\<^sub>E Roots)) (tuples Y) (tuples Y)" if \: "\ permutes Roots" and "Y \ Roots_ms" for \ Y proof (rule bij_betwI) have *: "(\f. restrict (\g. f (restrict (\ \ g) X)) (X \\<^sub>E Roots)) \ tuples Y \ tuples Y" if \: "\ permutes Roots" for \ proof fix f assume f: "f \ tuples Y" hence f': "f \ (X \\<^sub>E Roots) \\<^sub>E X" by (auto simp: tuples_def) define f' where "f' = (\g. f (restrict (\ \ g) X))" have "f' \ (X \\<^sub>E Roots) \ X" unfolding f'_def using f' bij_betw_apply[OF bij_betw_compose_left_perm_PiE[OF \, of X]] by blast hence "restrict f' (X \\<^sub>E Roots) \ (X \\<^sub>E Roots) \\<^sub>E X" by simp moreover have "image_mset (restrict f' (X \\<^sub>E Roots)) (mset_set (X \\<^sub>E Roots)) = Y" proof - have "image_mset (restrict f' (X \\<^sub>E Roots)) (mset_set (X \\<^sub>E Roots)) = image_mset f' (mset_set (X \\<^sub>E Roots))" by (intro image_mset_cong) auto also have "\ = image_mset f (image_mset (\g. restrict (\ \ g) X) (mset_set (X \\<^sub>E Roots)))" unfolding f'_def o_def multiset.map_comp by (simp add: o_def) also have "image_mset (\g. restrict (\ \ g) X) (mset_set (X \\<^sub>E Roots)) = mset_set (X \\<^sub>E Roots)" by (intro bij_betw_image_mset_set bij_betw_compose_left_perm_PiE \) also have "image_mset f \ = Y" using f by (simp add: tuples_def) finally show ?thesis . qed ultimately show "restrict f' (X \\<^sub>E Roots) \ tuples Y" by (auto simp: tuples_def) qed show "(\f. restrict (\g. f (restrict (\ \ g) X)) (X \\<^sub>E Roots)) \ tuples Y \ tuples Y" by (intro * \) show "(\f. restrict (\g. f (restrict (inv_into UNIV \ \ g) X)) (X \\<^sub>E Roots)) \ tuples Y \ tuples Y" by (intro * permutes_inv \) next have *: "(\g\X \\<^sub>E Roots. (\g\X \\<^sub>E Roots. f (restrict (\ \ g) X)) (restrict (inv_into UNIV \ \ g) X)) = f" (is "?lhs = _") if f: "f \ tuples Y" and \: "\ permutes Roots" for f \ proof fix g show "?lhs g = f g" proof (cases "g \ X \\<^sub>E Roots") case True have "restrict (\ \ restrict (inv_into UNIV \ \ g) X) X = g" using True by (intro ext) (auto simp: permutes_inverses[OF \]) thus ?thesis using True by (auto simp: permutes_in_image[OF permutes_inv[OF \]]) qed (use f in \auto simp: tuples_def\) qed show "(\g\X \\<^sub>E Roots. (\g\X \\<^sub>E Roots. f (restrict (\ \ g) X)) (restrict (inv_into UNIV \ \ g) X)) = f" if "f \ tuples Y" for f using *[OF that \] . show "(\g\X \\<^sub>E Roots. (\g\X \\<^sub>E Roots. f (restrict (inv_into UNIV \ \ g) X)) (restrict (\ \ g) X)) = f" if "f \ tuples Y" for f using *[OF that permutes_inv[OF \]] permutes_inv_inv[OF \] by simp qed text \ We show that the coefficients in the expanded new sum are rational -- again using the fundamental theorem of symmetric polynomials. \ define \' :: "complex multiset \ complex" where "\' = (\Y. \f\tuples Y. \g\X \\<^sub>E Roots. g (f g))" have "\' Y \ \" if Y: "Y \ Roots_ms" for Y proof - define Q :: "complex mpoly" where "Q = (\f\tuples Y. \g\X \\<^sub>E Roots. Var (unRoot (g (f g))))" have "insertion Root Q \ \" proof (rule symmetric_poly_of_roots_in_subring) show "ring_closed (\ :: complex set)" by standard auto then interpret ring_closed "\ :: complex set" . show "\m. coeff Q m \ \" by (auto simp: Q_def coeff_Var when_def intro!: sum_in_Rats coeff_prod_closed) next show "symmetric_mpoly {.. assume \: "\ permutes {..' where "\' = convert_perm (inv_into UNIV \)" have \': "\' permutes Roots" unfolding \'_def by (intro permutes_convert_perm permutes_inv \) have "mpoly_map_vars \ Q = (\f\tuples Y. \g\X \\<^sub>E Roots. Var (\ (unRoot (g (f g)))))" unfolding Q_def by (simp add: permutes_bij[OF \]) also have "\ = (\f\tuples Y. \g\X \\<^sub>E Roots. Var (unRoot (g (f (restrict (\' \ g) X)))))" proof (rule sum.cong, goal_cases) case (2 f) have f: "f \ (X \\<^sub>E Roots) \\<^sub>E X" using 2 by (auto simp: tuples_def) have "(\g\X \\<^sub>E Roots. Var (\ (unRoot (g (f g))))) = (\g\X \\<^sub>E Roots. Var (\ (unRoot (restrict (\' \ g) X (f (restrict (\' \ g) X))))))" using \' by (intro prod.reindex_bij_betw [symmetric] bij_betw_compose_left_perm_PiE) also have "\ = (\g\X \\<^sub>E Roots. Var (unRoot (g (f (restrict (\' \ g) X)))))" proof (intro prod.cong refl arg_cong[of _ _ Var]) fix g assume g: "g \ X \\<^sub>E Roots" have "restrict (\' \ g) X \ X \\<^sub>E Roots" using bij_betw_compose_left_perm_PiE[OF \', of X] g unfolding bij_betw_def by blast hence *: "f (restrict (\' \ g) X) \ X" by (rule PiE_mem[OF f]) hence **: "g (f (restrict (\' \ g) X)) \ Roots" by (rule PiE_mem[OF g]) have "unRoot (restrict (\' \ g) X (f (restrict (\' \ g) X))) = unRoot (Root (inv_into UNIV \ (unRoot (g (f (restrict (\' \ g) X))))))" using * ** by (subst \'_def) (auto simp: convert_perm_def) also have "inv_into UNIV \ (unRoot (g (f (restrict (\' \ g) X)))) \ {..]]) auto hence "unRoot (Root (inv_into UNIV \ (unRoot (g (f (restrict (\' \ g) X)))))) = inv_into UNIV \ (unRoot (g (f (restrict (\' \ g) X))))" by (intro unRoot_Root) auto also have "\ \ = unRoot (g (f (restrict (\' \ g) X)))" by (rule permutes_inverses[OF \]) finally show "\ (unRoot (restrict (\' \ g) X (f (restrict (\' \ g) X)))) = unRoot (g (f (restrict (\' \ g) X)))" . qed finally show ?case . qed simp_all also have "\ = (\x\tuples Y. \g\X \\<^sub>E Roots. Var (unRoot (g ((\g\X \\<^sub>E Roots. x (restrict (\' \ g) X)) g))))" by (intro sum.cong prod.cong refl) auto also have "\ = Q" unfolding Q_def by (rule sum.reindex_bij_betw[OF bij_betw_compose_perm]) (use \' Y in simp_all) finally show "mpoly_map_vars \ Q = Q" . qed next show "vars Q \ {.. 0" unfolding lc_def by auto thus "inverse (of_int lc) * (of_int lc :: complex) = 1" and "inverse (of_int lc) \ \" by auto have "rsquarefree (of_int_poly P :: complex poly)" using \squarefree (of_int_poly P :: complex poly)\ by (intro squarefree_imp_rsquarefree) hence "of_int_poly P = Polynomial.smult (of_int lc) (\x\Roots. [:-x, 1:])" unfolding lc_def of_int_hom.hom_lead_coeff[symmetric] Roots_def by (rule complex_poly_decompose_rsquarefree [symmetric]) also have "(\x\Roots. [:-x, 1:]) = (\iif\tuples Y. \g\X \\<^sub>E Roots. Root (unRoot (g (f g))))" by (simp add: Q_def insertion_sum insertion_prod) also have "\ = \' Y" unfolding \'_def by (intro sum.cong prod.cong refl Root_unRoot) (auto simp: tuples_def) finally show ?thesis . qed hence "\Y\Roots_ms. \x. \' Y = of_rat x" by (auto elim!: Rats_cases) then obtain \'' :: "complex multiset \ rat" where \'': "\Y. Y \ Roots_ms \ \' Y = of_rat (\'' Y)" by metis text \ We again collect all the terms that happen to have equal exponents and call their coefficients \\''\: \ define \''' :: "complex \ rat" where "\''' = (\\. \Y\Roots_ms. (\'' Y when \\<^sub>#Y = \))" have supp_\''': "{x. \''' x \ 0} \ sum_mset ` Roots_ms" by (auto simp: \'''_def when_def elim!: sum.not_neutral_contains_not_neutral split: if_splits) text \ We again start with the sum that we now to be zero and multiply it with all the sums that can be obtained with different choices for the roots. \ have "0 = (\x\X. \ x * exp x)" using sum0 .. also have "\ = (\x\X. restrict \ X x * exp x)" by (intro sum.cong) auto also have "\ dvd (\f \ X \\<^sub>E Roots. \x\X. f x * exp x)" by (rule dvd_prodI) (use \\ ` X \ Roots\ in \auto simp: id_def\) also have "\ = (\f\(X \\<^sub>E Roots) \\<^sub>E X. \g\X \\<^sub>E Roots. g (f g) * exp (f g))" by (rule prod_sum_PiE) auto also have "\ = (\f\(X \\<^sub>E Roots) \\<^sub>E X. (\g\X \\<^sub>E Roots. g (f g)) * exp (\g\X \\<^sub>E Roots. f g))" by (simp add: prod.distrib exp_sum) also have "\ = (\(Y,f)\Sigma Roots_ms tuples. (\g\X \\<^sub>E Roots. g (f g)) * exp (\g\X \\<^sub>E Roots. f g))" by (intro sum.reindex_bij_witness[of _ snd "\f. (image_mset f (mset_set (X \\<^sub>E Roots)), f)"]) (auto simp: Roots_ms_def tuples_def) also have "\ = (\(Y,f)\Sigma Roots_ms tuples. (\g\X \\<^sub>E Roots. g (f g)) * exp (\\<^sub>#Y))" by (intro sum.cong) (auto simp: tuples_def sum_unfold_sum_mset) also have "\ = (\Y\Roots_ms. \' Y * exp (\\<^sub>#Y))" unfolding \'_def sum_distrib_right by (rule sum.Sigma [symmetric]) auto also have "\ = (\Y\Roots_ms. of_rat (\'' Y) * exp (\\<^sub>#Y))" by (intro sum.cong) (auto simp: \'') also have "\ = (\Y\Roots_ms. Sum_any (\\. of_rat (\'' Y when \\<^sub># Y = \) * exp \))" proof (rule sum.cong, goal_cases) case (2 Y) have "Sum_any (\\. of_rat (\'' Y when \\<^sub># Y = \) * exp \) = (\\\{\\<^sub># Y}. of_rat (\'' Y when \\<^sub># Y = \) * exp \)" by (intro Sum_any.expand_superset) auto thus ?case by simp qed auto also have "\ = Sum_any (\\. of_rat (\''' \) * exp \)" unfolding \'''_def of_rat_sum sum_distrib_right by (subst Sum_any_sum_swap) auto also have "\ = (\\ | \''' \ \ 0. of_rat (\''' \) * exp \)" by (intro Sum_any.expand_superset finite_subset[OF supp_\''']) auto finally have "(\\ | \''' \ \ 0. of_rat (\''' \) * exp \) = 0" by auto text \ We are now in the situation of our previous version of the theorem and can apply it to find that all the coefficients are zero. \ have "{\. \''' \ \ 0} = {}" proof (rule Hermite_Lindemann_aux3) show "finite {\. \''' \ \ 0}" by (rule finite_subset[OF supp_\''']) auto next show "(\\ | \''' \ \ 0. of_rat (\''' \) * exp \) = 0" by fact next fix \ assume \: "\ \ {\. \''' \ \ 0}" then obtain Y where Y: "Y \ Roots_ms" "\ = sum_mset Y" using supp_\''' by auto thus "algebraic \" using alg1 by (auto simp: Roots_ms_def) qed auto text \ However, similarly to before, we can show that the coefficient corresponding to the term with the lexicographically greatest exponent (which is obtained by picking the term with the lexicographically greatest term in each of the factors of our big product) is non-zero. \ moreover have "\\. \''' \ \ 0" proof - define \_max where "\_max = complex_lex.Max X" have [simp]: "\_max \ X" unfolding \_max_def using \X \ {}\ by (intro complex_lex.Max_in) auto define Y_max :: "complex multiset" where "Y_max = replicate_mset (n ^ m) \_max" define f_max where "f_max = restrict (\_. \_max) (X \\<^sub>E Roots)" have [simp]: "Y_max \ Roots_ms" by (auto simp: Y_max_def Roots_ms_def) have "tuples Y_max = {f_max}" proof safe have "image_mset (\_\X \\<^sub>E Roots. \_max) (mset_set (X \\<^sub>E Roots)) = image_mset (\_. \_max) (mset_set (X \\<^sub>E Roots))" by (intro image_mset_cong) auto thus "f_max \ tuples Y_max" by (auto simp: f_max_def tuples_def Y_max_def image_mset_const_eq) next fix f assume "f \ tuples Y_max" hence f: "f \ (X \\<^sub>E Roots) \\<^sub>E X" "image_mset f (mset_set (X \\<^sub>E Roots)) = Y_max" by (auto simp: tuples_def) hence "\g \# mset_set (X \\<^sub>E Roots). f g = \_max" by (intro image_mset_eq_replicate_msetD[where n = "n ^ m"]) (auto simp: Y_max_def) thus "f = f_max" using f by (auto simp: Y_max_def fun_eq_iff f_max_def) qed have "\''' (of_nat (n ^ m) * \_max) = (\Y\Roots_ms. \'' Y when \\<^sub># Y = of_nat (n ^ m) * \_max)" unfolding \'''_def Roots_ms_def .. also have "\\<^sub># Y \ of_nat n ^ m * \_max" if "Y \ Roots_ms" "Y \ Y_max" for Y proof - have "\set_mset Y \ {\_max}" using set_mset_subset_singletonD[of Y "\_max"] that by (auto simp: Roots_ms_def Y_max_def split: if_splits) then obtain y where y: "y \# Y" "y \ \_max" by auto have "y \ X" "set_mset (Y - {#y#}) \ X" using y that by (auto simp: Roots_ms_def dest: in_diffD) hence "y \\<^sub>\ \_max" using y unfolding \_max_def by (intro complex_lex.Max_ge) auto with y have "y <\<^sub>\ \_max" by auto have *: "Y = {#y#} + (Y - {#y#})" using y by simp have "sum_mset Y = y + sum_mset (Y - {#y#})" by (subst *) auto also have "\ <\<^sub>\ \_max + sum_mset (Y - {#y#})" by (intro complex_lex.add_strict_right_mono) fact also have "\ \\<^sub>\ \_max + sum_mset (replicate_mset (n ^ m - 1) \_max)" unfolding \_max_def using that y \set_mset (Y - {#y#}) \ X\ by (intro complex_lex.add_left_mono sum_mset_mono_complex_lex rel_mset_replicate_mset_right complex_lex.Max_ge) (auto simp: Roots_ms_def size_Diff_singleton) also have "\ = of_nat (Suc (n ^ m - 1)) * \_max" by (simp add: algebra_simps) also have "Suc (n ^ m - 1) = n ^ m" using \n > 0\ by simp finally show ?thesis by simp qed hence "(\Y\Roots_ms. \'' Y when \\<^sub># Y = of_nat (n ^ m) * \_max) = (\Y\{Y_max}. \'' Y when \\<^sub># Y = of_nat (n ^ m) * \_max)" by (intro sum.mono_neutral_right ballI) auto also have "\ = \'' Y_max" by (auto simp: when_def Y_max_def) also have "of_rat \ = \' Y_max" using \''[of Y_max] by auto also have "\ = (\g\X \\<^sub>E Roots. g (f_max g))" by (auto simp: \'_def \tuples Y_max = {f_max}\) also have "\ = (\g\X \\<^sub>E Roots. g \_max)" by (intro prod.cong) (auto simp: f_max_def) also have "\ \ 0" using \0 \ Roots\ \\_max \ X\ by (intro prod_nonzeroI) (metis PiE_mem) finally show ?thesis by blast qed ultimately show False by blast qed subsection \The final theorem\ text \ We now additionally allow some of the $\beta_i$ to be zero: \ lemma Hermite_Lindemann': fixes \ :: "complex \ complex" assumes "finite X" assumes "\x. x \ X \ algebraic x" assumes "\x. x \ X \ algebraic (\ x)" assumes "(\x\X. \ x * exp x) = 0" shows "\x\X. \ x = 0" proof - have "{x\X. \ x \ 0} = {}" proof (rule Hermite_Lindemann_aux4) have "(\x | x \ X \ \ x \ 0. \ x * exp x) = (\x\X. \ x * exp x)" by (intro sum.mono_neutral_left assms(1)) auto also have "\ = 0" by fact finally show "(\x | x \ X \ \ x \ 0. \ x * exp x) = 0" . qed (use assms in auto) thus ?thesis by blast qed text \ Lastly, we switch to indexed summation in order to obtain a version of the theorem that is somewhat nicer to use: \ theorem Hermite_Lindemann: fixes \ \ :: "'a \ complex" assumes "finite I" assumes "\x. x \ I \ algebraic (\ x)" assumes "\x. x \ I \ algebraic (\ x)" assumes "inj_on \ I" assumes "(\x\I. \ x * exp (\ x)) = 0" shows "\x\I. \ x = 0" proof - define f where "f = inv_into I \" have [simp]: "f (\ x) = x" if "x \ I" for x using that by (auto simp: f_def inv_into_f_f[OF assms(4)]) have "\x\\`I. (\ \ f) x = 0" proof (rule Hermite_Lindemann') have "0 = (\x\I. \ x * exp (\ x))" using assms(5) .. also have "\ = (\x\I. (\ \ f) (\ x) * exp (\ x))" by (intro sum.cong) auto also have "\ = (\x\\`I. (\ \ f) x * exp x)" using assms(4) by (subst sum.reindex) auto finally show "(\x\\ ` I. (\ \ f) x * exp x) = 0" .. qed (use assms in auto) thus ?thesis by auto qed text \ The following version using lists instead of sequences is even more convenient to use in practice: \ corollary Hermite_Lindemann_list: fixes xs :: "(complex \ complex) list" assumes alg: "\(x,y)\set xs. algebraic x \ algebraic y" assumes distinct: "distinct (map snd xs)" assumes sum0: "(\(c,\)\xs. c * exp \) = 0" shows "\c\(fst ` set xs). c = 0" proof - define n where "n = length xs" have *: "\i\{..i. map snd xs ! i) {.. inj_on (\i. snd (xs ! i)) {..i. snd (xs ! i)) {..(c,\)\xs. c * exp \)" using sum0 .. also have "\ = (\i = 0" .. next fix i assume i: "i \ {.. set xs" by (auto simp: n_def) with alg show "algebraic (fst (xs ! i))" "algebraic (snd (xs ! i))" by blast+ qed auto show ?thesis proof (intro ballI, elim imageE) fix c x assume cx: "c = fst x" "x \ set xs" then obtain i where "i \ {..The traditional formulation of the theorem\ text \ What we proved above was actually Baker's reformulation of the theorem. Thus, we now also derive the original one, which uses linear independence and algebraic independence. It states that if $\alpha_1, \ldots, \alpha_n$ are algebraic numbers that are linearly independent over \\\, then $e^{\alpha_1}, \ldots, e^{\alpha_n}$ are algebraically independent over \\\. \ text \ Linear independence over the integers is just independence of a set of complex numbers when viewing the complex numbers as a \\\-module. \ definition linearly_independent_over_int :: "'a :: field_char_0 set \ bool" where "linearly_independent_over_int = module.independent (\r x. of_int r * x)" text \ Algebraic independence over the rationals means that the given set \X\ of numbers fulfils no non-trivial polynomial equation with rational coefficients, i.e. there is no non-zero multivariate polynomial with rational coefficients that, when inserting the numbers from \X\, becomes zero. Note that we could easily replace `rational coefficients' with `algebraic coefficients' here and the proof would still go through without any modifications. \ definition algebraically_independent_over_rat :: "nat \ (nat \ 'a :: field_char_0) \ bool" where "algebraically_independent_over_rat n a \ (\p. vars p \ {.. (\m. coeff p m \ \) \ insertion a p = 0 \ p = 0)" corollary Hermite_Lindemann_original: fixes n :: nat and \ :: "nat \ complex" assumes "inj_on \ {..i. i < n \ algebraic (\ i)" assumes "linearly_independent_over_int (\ ` {..i. exp (\ i))" unfolding algebraically_independent_over_rat_def proof safe fix p assume p: "vars p \ {..m. coeff p m \ \" "insertion (\i. exp (\ i)) p = 0" define \' where "\' = (\m. \i i)" define I where "I = {m. coeff p m \ 0}" have lookup_eq_0: "lookup m i = 0" if "m \ I" "i \ {.. vars p" using that coeff_notin_vars[of m p] by (auto simp: I_def) thus "lookup m i = 0" using in_keys_iff[of i m] that p(1) by blast qed have "\x\I. coeff p x = 0" proof (rule Hermite_Lindemann) show "finite I" by (auto simp: I_def) next show "algebraic (\' m)" if "m \ I" for m unfolding \'_def using assms(2) by fastforce next show "algebraic (coeff p m)" if "m \ I" for m unfolding \'_def using p(2) by blast next show "inj_on \' I" proof fix m1 m2 assume m12: "m1 \ I" "m2 \ I" "\' m1 = \' m2" define lu :: "(nat \\<^sub>0 nat) \ nat \ int" where "lu = (\m i. int (lookup m i))" interpret int: Modules.module "\r x. of_int r * (x :: complex)" by standard (auto simp: algebra_simps of_rat_mult of_rat_add) define idx where "idx = inv_into {.." have "lu m1 i = lu m2 i" if "i < n" for i proof - have "lu m1 (idx (\ i)) - lu m2 (idx (\ i)) = 0" proof (rule int.independentD) show "int.independent (\ ` {..x\\`{..i i)) - lu m2 (idx (\ i))) * \ i)" using assms(1) by (subst sum.reindex) auto also have "\ = (\i i)" by (intro sum.cong) (auto simp: idx_def inv_into_f_f[OF assms(1)]) also have "\ = 0" using m12 by (simp add: \'_def ring_distribs of_rat_diff sum_subtractf lu_def) finally show "(\x\\`{..'_def ring_distribs of_rat_diff sum_subtractf lu_def) qed (use that in auto) thus ?thesis using that by (auto simp: idx_def inv_into_f_f[OF assms(1)]) qed hence "lookup m1 i = lookup m2 i" for i using m12 by (cases "i < n") (auto simp: lu_def lookup_eq_0) thus "m1 = m2" by (rule poly_mapping_eqI) qed next have "0 = insertion (\i. exp (\ i)) p" using p(3) .. also have "\ = (\m\I. coeff p m * Prod_any (\i. exp (\ i) ^ lookup m i))" unfolding insertion_altdef by (rule Sum_any.expand_superset) (auto simp: I_def) also have "\ = (\m\I. coeff p m * exp (\' m))" proof (intro sum.cong, goal_cases) case (2 m) have "Prod_any (\i. exp (\ i) ^ lookup m i) = (\i i) ^ lookup m i)" using 2 lookup_eq_0[of m] by (intro Prod_any.expand_superset; force) also have "\ = exp (\' m)" by (simp add: exp_sum exp_of_nat_mult \'_def) finally show ?case by simp qed simp_all finally show "(\m\I. coeff p m * exp (\' m)) = 0" .. qed thus "p = 0" by (intro mpoly_eqI) (auto simp: I_def) qed subsection \Simple corollaries\ text \ Now, we derive all the usual obvious corollaries of the theorem in the obvious way. First, the exponential of a non-zero algebraic number is transcendental. \ corollary algebraic_exp_complex_iff: assumes "algebraic x" shows "algebraic (exp x :: complex) \ x = 0" using Hermite_Lindemann_list[of "[(1, x), (-exp x, 0)]"] assms by auto text \ More generally, any sum of exponentials with algebraic coefficients and exponents is transcendental if the exponents are all distinct and non-zero and at least one coefficient is non-zero. \ corollary sum_of_exp_transcendentalI: fixes xs :: "(complex \ complex) list" assumes "\(x,y)\set xs. algebraic x \ algebraic y \ y \ 0" assumes "\x\fst`set xs. x \ 0" assumes distinct: "distinct (map snd xs)" shows "\algebraic (\(c,\)\xs. c * exp \)" proof define S where "S = (\(c,\)\xs. c * exp \)" assume S: "algebraic S" have "\c\fst`set ((-S,0) # xs). c = 0" proof (rule Hermite_Lindemann_list) show "(\(c, \)\(- S, 0) # xs. c * exp \) = 0" by (auto simp: S_def) qed (use S assms in auto) with assms(2) show False by auto qed text \ Any complex logarithm of an algebraic number other than 1 is transcendental (no matter which branch cut). \ corollary transcendental_complex_logarithm: assumes "algebraic x" "exp y = (x :: complex)" "x \ 1" shows "\algebraic y" using algebraic_exp_complex_iff[of y] assms by auto text \ In particular, this holds for the standard branch of the logarithm. \ corollary transcendental_Ln: assumes "algebraic x" "x \ 0" "x \ 1" shows "\algebraic (Ln x)" by (rule transcendental_complex_logarithm) (use assms in auto) text \ The transcendence of \e\ and \\\, which I have already formalised directly in other AFP entries, now follows as a simple corollary. \ corollary exp_1_complex_transcendental: "\algebraic (exp 1 :: complex)" by (subst algebraic_exp_complex_iff) auto corollary pi_transcendental: "\algebraic pi" proof - have "\algebraic (\ * pi)" by (rule transcendental_complex_logarithm[of "-1"]) auto thus ?thesis by simp qed subsection \Transcendence of the trigonometric and hyperbolic functions\ text \ In a similar fashion, we can also prove the transcendence of all the trigonometric and hyperbolic functions such as $\sin$, $\tan$, $\sinh$, $\arcsin$, etc. \ lemma transcendental_sinh: assumes "algebraic z" "z \ 0" shows "\algebraic (sinh z :: complex)" proof - have "\algebraic (\(a,b)\[(1/2, z), (-1/2, -z)]. a * exp b)" using assms by (intro sum_of_exp_transcendentalI) auto also have "(\(a,b)\[(1/2, z), (-1/2, -z)]. a * exp b) = sinh z" by (simp add: sinh_def field_simps scaleR_conv_of_real) finally show ?thesis . qed lemma transcendental_cosh: assumes "algebraic z" "z \ 0" shows "\algebraic (cosh z :: complex)" proof - have "\algebraic (\(a,b)\[(1/2, z), (1/2, -z)]. a * exp b)" using assms by (intro sum_of_exp_transcendentalI) auto also have "(\(a,b)\[(1/2, z), (1/2, -z)]. a * exp b) = cosh z" by (simp add: cosh_def field_simps scaleR_conv_of_real) finally show ?thesis . qed lemma transcendental_sin: assumes "algebraic z" "z \ 0" shows "\algebraic (sin z :: complex)" unfolding sin_conv_sinh using transcendental_sinh[of "\ * z"] assms by simp lemma transcendental_cos: assumes "algebraic z" "z \ 0" shows "\algebraic (cos z :: complex)" unfolding cos_conv_cosh using transcendental_cosh[of "\ * z"] assms by simp (* TODO: Move? *) lemma tan_square_neq_neg1: "tan (z :: complex) ^ 2 \ -1" proof assume "tan z ^ 2 = -1" hence "sin z ^ 2 = -(cos z ^ 2)" by (auto simp: tan_def divide_simps split: if_splits) also have "cos z ^ 2 = 1 - sin z ^ 2" by (simp add: cos_squared_eq) finally show False by simp qed lemma transcendental_tan: assumes "algebraic z" "z \ 0" shows "\algebraic (tan z :: complex)" proof assume "algebraic (tan z)" have nz1: "real_of_int n + 1 / 2 \ 0" for n proof - have "real_of_int (2 * n + 1) / real_of_int 2 \ \" by (intro fraction_not_in_ints) auto also have "real_of_int (2 * n + 1) / real_of_int 2 = real_of_int n + 1 / 2" by simp finally show "\ \ 0" by auto qed have nz2: "1 + tan z ^ 2 \ 0" using tan_square_neq_neg1[of z] by (subst add_eq_0_iff) have nz3: "cos z \ 0" proof assume "cos z = 0" then obtain n where "z = complex_of_real (real_of_int n * pi) + complex_of_real pi / 2" by (subst (asm) cos_eq_0) blast also have "\ = complex_of_real ((real_of_int n + 1 / 2) * pi)" by (simp add: ring_distribs) also have "algebraic \ \ algebraic ((real_of_int n + 1 / 2) * pi)" by (rule algebraic_of_real_iff) also have "\algebraic ((real_of_int n + 1 / 2) * pi)" using nz1[of n] transcendental_pi by simp finally show False using assms(1) by contradiction qed from nz3 have *: "sin z ^ 2 = tan z ^ 2 * cos z ^ 2" by (simp add: tan_def field_simps) also have "cos z ^ 2 = 1 - sin z ^ 2" by (simp add: cos_squared_eq) finally have "sin z ^ 2 * (1 + tan z ^ 2) = tan z ^ 2" by (simp add: algebra_simps) hence "sin z ^ 2 = tan z ^ 2 / (1 + tan z ^ 2)" using nz2 by (simp add: field_simps) also have "algebraic (tan z ^ 2 / (1 + tan z ^ 2))" using \algebraic (tan z)\ by auto finally have "algebraic (sin z ^ 2)" . hence "algebraic (sin z)" by simp thus False using transcendental_sin[OF \algebraic z\ \z \ 0\] by contradiction qed lemma transcendental_cot: assumes "algebraic z" "z \ 0" shows "\algebraic (cot z :: complex)" proof - have "\algebraic (tan z)" by (rule transcendental_tan) fact+ also have "algebraic (tan z) \ algebraic (inverse (tan z))" by simp also have "inverse (tan z) = cot z" by (simp add: cot_def tan_def) finally show ?thesis . qed lemma transcendental_tanh: assumes "algebraic z" "z \ 0" "cosh z \ 0" shows "\algebraic (tanh z :: complex)" using transcendental_tan[of "\ * z"] assms unfolding tanh_conv_tan by simp lemma transcendental_Arcsin: assumes "algebraic z" "z \ 0" shows "\algebraic (Arcsin z)" proof - have "\ * z + csqrt (1 - z\<^sup>2) \ 0" using Arcsin_body_lemma by blast moreover have "\ * z + csqrt (1 - z\<^sup>2) \ 1" proof assume "\ * z + csqrt (1 - z\<^sup>2) = 1" hence "Arcsin z = 0" by (simp add: Arcsin_def) hence "sin (Arcsin z) = 0" by (simp only: sin_zero) also have "sin (Arcsin z) = z" by simp finally show False using \z \ 0\ by simp qed ultimately have "\ algebraic (Ln (\ * z + csqrt (1 - z\<^sup>2)))" using assms by (intro transcendental_Ln) auto thus ?thesis by (simp add: Arcsin_def) qed lemma transcendental_Arccos: assumes "algebraic z" "z \ 1" shows "\algebraic (Arccos z)" proof - have "z + \ * csqrt (1 - z\<^sup>2) \ 0" using Arccos_body_lemma by blast moreover have "z + \ * csqrt (1 - z\<^sup>2) \ 1" proof assume "z + \ * csqrt (1 - z\<^sup>2) = 1" hence "Arccos z = 0" by (simp add: Arccos_def) hence "cos (Arccos z) = 1" by (simp only: cos_zero) also have "cos (Arccos z) = z" by simp finally show False using \z \ 1\ by simp qed ultimately have "\ algebraic (Ln (z + \ * csqrt (1 - z\<^sup>2)))" using assms by (intro transcendental_Ln) auto thus ?thesis by (simp add: Arccos_def) qed lemma transcendental_Arctan: assumes "algebraic z" "z \ {0, \, -\}" shows "\algebraic (Arctan z)" proof - have "\ * z \ 1" "1 + \ * z \ 0" using assms(2) by (auto simp: complex_eq_iff) hence "\ algebraic (Ln ((1 - \ * z) / (1 + \ * z)))" using assms by (intro transcendental_Ln) auto thus ?thesis by (simp add: Arctan_def) qed end \ No newline at end of file diff --git a/thys/Zeta_3_Irrational/Zeta_3_Irrational.thy b/thys/Zeta_3_Irrational/Zeta_3_Irrational.thy --- a/thys/Zeta_3_Irrational/Zeta_3_Irrational.thy +++ b/thys/Zeta_3_Irrational/Zeta_3_Irrational.thy @@ -1,2089 +1,2092 @@ (* File: Zeta_3_Irrational.thy Author: Manuel Eberl, TU München *) section \The Irrationality of $\zeta(3)$\ theory Zeta_3_Irrational imports "E_Transcendental.E_Transcendental" "Prime_Number_Theorem.Prime_Number_Theorem" "Prime_Distribution_Elementary.PNT_Consequences" begin +hide_const (open) UnivPoly.coeff UnivPoly.up_ring.monom +hide_const (open) Module.smult Coset.order + text \ Ap\'{e}ry's original proof of the irrationality of $\zeta(3)$ contained several tricky identities of sums involving binomial coefficients that are difficult to prove. I instead follow Beukers's proof, which consists of several fairly straightforward manipulations of integrals with absolutely no caveats. Both Ap\'{e}ry and Beukers make use of an asymptotic upper bound on $\text{lcm}\{1\ldots n\}$ -- namely $\text{lcm}\{1\ldots n\} \in o(c^n)$ for any $c > e$, which is a consequence of the Prime Number Theorem (which, fortunately, is available in the \emph{Archive of Formal Proofs}~\<^cite>\"afp_primes1" and "afp_primes2"\). I follow the concise presentation of Beukers's proof in Filaseta's lecture notes~\<^cite>\"filaseta"\, which turned out to be very amenable to formalisation. There is another earlier formalisation of the irrationality of $\zeta(3)$ by Mahboubi\ \emph{et al.}~\<^cite>\"mahboubi"\, who followed Ap\'{e}ry's original proof, but were ultimately forced to find a more elementary way to prove the asymptotics of $\text{lcm}\{1\ldots n\}$ than the Prime Number Theorem. First, we will require some auxiliary material before we get started with the actual proof. \ (* TODO: A lot of this (if not all of it) should really be in the library *) subsection \Auxiliary facts about polynomials\ lemma degree_higher_pderiv: "degree ((pderiv ^^ n) p) = degree p - n" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" by (induction n arbitrary: p) (auto simp: degree_pderiv) lemma pcompose_power_left: "pcompose (p ^ n) q = pcompose p q ^ n" by (induction n) (auto simp: pcompose_mult pcompose_1) lemma pderiv_sum: "pderiv (\x\A. f x) = (\x\A. pderiv (f x))" by (induction A rule: infinite_finite_induct) (auto simp: pderiv_add) lemma higher_pderiv_minus: "(pderiv ^^ n) (-p :: 'a :: idom poly) = -(pderiv ^^ n) p" by (induction n) (auto simp: pderiv_minus) 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 pderiv_monom: "pderiv (monom c n) = monom (of_nat n * c) (n - 1)" by (simp add: monom_altdef pderiv_smult pderiv_power pderiv_pCons mult_ac) lemma higher_pderiv_monom: "k \ n \ (pderiv ^^ k) (monom c n) = monom (of_nat (pochhammer (n - k + 1) k) * c) (n - k)" by (induction k) (auto simp: pderiv_monom pochhammer_rec Suc_diff_le Suc_diff_Suc mult_ac) lemma higher_pderiv_mult: "(pderiv ^^ n) (p * q) = (\k\n. Polynomial.smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (n - k)) q))" proof (induction n) case (Suc n) have eq: "(Suc n choose k) = (n choose k) + (n choose (k-1))" if "k > 0" for k using that by (cases k) auto have "(pderiv ^^ Suc n) (p * q) = (\k\n. smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) + (\k\n. smult (of_nat (n choose k)) ((pderiv ^^ Suc k) p * (pderiv ^^ (n - k)) q))" by (simp add: Suc pderiv_sum pderiv_smult pderiv_mult sum.distrib smult_add_right algebra_simps Suc_diff_le) also have "(\k\n. smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) = (\k\insert 0 {1..n}. smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q))" by (intro sum.cong) auto also have "\ = (\k=1..n. smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) + p * (pderiv ^^ Suc n) q" by (subst sum.insert) (auto simp: add_ac) also have "(\k\n. smult (of_nat (n choose k)) ((pderiv ^^ Suc k) p * (pderiv ^^ (n - k)) q)) = (\k=1..n+1. smult (of_nat (n choose (k-1))) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q))" by (intro sum.reindex_bij_witness[of _ "\k. k - 1" Suc]) auto also have "\ = (\k\insert (n+1) {1..n}. smult (of_nat (n choose (k-1))) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q))" by (intro sum.cong) auto also have "\ = (\k=1..n. smult (of_nat (n choose (k-1))) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) + (pderiv ^^ Suc n) p * q" by (subst sum.insert) (auto simp: add_ac) also have "(\k=1..n. smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) + p * (pderiv ^^ Suc n) q + \ = (\k=1..n. smult (of_nat (Suc n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) + p * (pderiv ^^ Suc n) q + (pderiv ^^ Suc n) p * q" by (simp add: sum.distrib algebra_simps smult_add_right eq smult_add_left) also have "\ = (\k\{1..n}\{0,Suc n}. smult (of_nat (Suc n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q))" by (subst sum.union_disjoint) (auto simp: algebra_simps) also have "{1..n}\{0,Suc n} = {..Suc n}" by auto finally show ?case . qed auto subsection \Auxiliary facts about integrals\ theorem (in pair_sigma_finite) Fubini_set_integrable: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f[measurable]: "set_borel_measurable (M1 \\<^sub>M M2) (A \ B) f" and integ1: "set_integrable M1 A (\x. \y\B. norm (f (x, y)) \M2)" and integ2: "AE x\A in M1. set_integrable M2 B (\y. f (x, y))" shows "set_integrable (M1 \\<^sub>M M2) (A \ B) f" unfolding set_integrable_def proof (rule Fubini_integrable) note integ1 also have "set_integrable M1 A (\x. \y\B. norm (f (x, y)) \M2) \ integrable M1 (\x. LINT y|M2. norm (indicat_real (A \ B) (x, y) *\<^sub>R f (x, y)))" unfolding set_integrable_def by (intro Bochner_Integration.integrable_cong) (auto simp: indicator_def set_lebesgue_integral_def) finally show \ . next from integ2 show "AE x in M1. integrable M2 (\y. indicat_real (A \ B) (x, y) *\<^sub>R f (x, y))" proof eventually_elim case (elim x) show "integrable M2 (\y. indicat_real (A \ B) (x, y) *\<^sub>R f (x, y))" proof (cases "x \ A") case True with elim have "set_integrable M2 B (\y. f (x, y))" by simp also have "?this \ ?thesis" unfolding set_integrable_def using True by (intro Bochner_Integration.integrable_cong) (auto simp: indicator_def) finally show ?thesis . qed auto qed qed (insert assms, auto simp: set_borel_measurable_def) lemma (in pair_sigma_finite) set_integral_fst': fixes f :: "_ \ 'c :: {second_countable_topology, banach}" assumes "set_integrable (M1 \\<^sub>M M2) (A \ B) f" shows "set_lebesgue_integral (M1 \\<^sub>M M2) (A \ B) f = (\x\A. (\y\B. f (x, y) \M2) \M1)" proof - have "set_lebesgue_integral (M1 \\<^sub>M M2) (A \ B) f = (\z. indicator (A \ B) z *\<^sub>R f z \(M1 \\<^sub>M M2))" by (simp add: set_lebesgue_integral_def) also have "\ = (\x. \y. indicator (A \ B) (x,y) *\<^sub>R f (x,y) \M2 \M1)" using assms by (subst integral_fst' [symmetric]) (auto simp: set_integrable_def) also have "\ = (\x\A. (\y\B. f (x,y) \M2) \M1)" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_cong refl) (auto simp: indicator_def) finally show ?thesis . qed lemma (in pair_sigma_finite) set_integral_snd: fixes f :: "_ \ 'c :: {second_countable_topology, banach}" assumes "set_integrable (M1 \\<^sub>M M2) (A \ B) f" shows "set_lebesgue_integral (M1 \\<^sub>M M2) (A \ B) f = (\y\B. (\x\A. f (x, y) \M1) \M2)" proof - have "set_lebesgue_integral (M1 \\<^sub>M M2) (A \ B) f = (\z. indicator (A \ B) z *\<^sub>R f z \(M1 \\<^sub>M M2))" by (simp add: set_lebesgue_integral_def) also have "\ = (\y. \x. indicator (A \ B) (x,y) *\<^sub>R f (x,y) \M1 \M2)" using assms by (subst integral_snd) (auto simp: set_integrable_def case_prod_unfold) also have "\ = (\y\B. (\x\A. f (x,y) \M1) \M2)" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_cong refl) (auto simp: indicator_def) finally show ?thesis . qed proposition (in pair_sigma_finite) Fubini_set_integral: fixes f :: "_ \ _ \ _ :: {banach, second_countable_topology}" assumes f: "set_integrable (M1 \\<^sub>M M2) (A \ B) (case_prod f)" shows "(\y\B. (\x\A. f x y \M1) \M2) = (\x\A. (\y\B. f x y \M2) \M1)" proof - have "(\y\B. (\x\A. f x y \M1) \M2) = (\y. (\x. indicator (A \ B) (x, y) *\<^sub>R f x y \M1) \M2)" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_cong) (auto simp: indicator_def) also have "\ = (\x. (\y. indicator (A \ B) (x, y) *\<^sub>R f x y \M2) \M1)" using assms by (intro Fubini_integral) (auto simp: set_integrable_def case_prod_unfold) also have "\ = (\x\A. (\y\B. f x y \M2) \M1)" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_cong) (auto simp: indicator_def) finally show ?thesis . qed lemma (in pair_sigma_finite) nn_integral_swap: assumes [measurable]: "f \ borel_measurable (M1 \\<^sub>M M2)" shows "(\\<^sup>+x. f x \(M1 \\<^sub>M M2)) = (\\<^sup>+(y,x). f (x,y) \(M2 \\<^sub>M M1))" by (subst distr_pair_swap, subst nn_integral_distr) (auto simp: case_prod_unfold) lemma set_integrable_bound: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ 'c::{banach, second_countable_topology}" shows "set_integrable M A f \ set_borel_measurable M A g \ (AE x in M. x \ A \ norm (g x) \ norm (f x)) \ set_integrable M A g" unfolding set_integrable_def apply (erule Bochner_Integration.integrable_bound) apply (simp add: set_borel_measurable_def) apply (erule eventually_mono) apply (auto simp: indicator_def) done lemma set_integrableI_nonneg: fixes f :: "'a \ real" assumes "set_borel_measurable M A f" assumes "AE x in M. x \ A \ 0 \ f x" "(\\<^sup>+x\A. f x \M) < \" shows "set_integrable M A f" unfolding set_integrable_def proof (rule integrableI_nonneg) from assms show "(\x. indicator A x *\<^sub>R f x) \ borel_measurable M" by (simp add: set_borel_measurable_def) from assms(2) show "AE x in M. 0 \ indicat_real A x *\<^sub>R f x" by eventually_elim (auto simp: indicator_def) have "(\\<^sup>+x. ennreal (indicator A x *\<^sub>R f x) \M) = (\\<^sup>+x\A. f x \M)" by (intro nn_integral_cong) (auto simp: indicator_def) also have "\ < \" by fact finally show "(\\<^sup>+x. ennreal (indicator A x *\<^sub>R f x) \M) < \" . qed lemma set_integral_eq_nn_integral: assumes "set_borel_measurable M A f" assumes "set_nn_integral M A f = ennreal x" "x \ 0" assumes "AE x in M. x \ A \ f x \ 0" shows "set_integrable M A f" and "set_lebesgue_integral M A f = x" proof - have eq: "(\x. (if x \ A then 1 else 0) * f x) = (\x. if x \ A then f x else 0)" "(\x. if x \ A then ennreal (f x) else 0) = (\x. ennreal (f x) * (if x \ A then 1 else 0))" "(\x. ennreal (f x * (if x \ A then 1 else 0))) = (\x. ennreal (f x) * (if x \ A then 1 else 0))" by auto from assms show *: "set_integrable M A f" by (intro set_integrableI_nonneg) auto have "set_lebesgue_integral M A f = enn2real (set_nn_integral M A f)" unfolding set_lebesgue_integral_def using assms(1,4) * eq by (subst integral_eq_nn_integral) (auto intro!: nn_integral_cong simp: indicator_def of_bool_def set_integrable_def mult_ac) also have "\ = x" using assms by simp finally show "set_lebesgue_integral M A f = x" . qed lemma set_integral_0 [simp, intro]: "set_integrable M A (\y. 0)" by (simp add: set_integrable_def) lemma set_integrable_sum: fixes f :: "_ \ _ \ _ :: {banach, second_countable_topology}" assumes "finite B" assumes "\x. x \ B \ set_integrable M A (f x)" shows "set_integrable M A (\y. \x\B. f x y)" using assms by (induction rule: finite_induct) (auto intro!: set_integral_add) lemma set_integral_sum: fixes f :: "_ \ _ \ _ :: {banach, second_countable_topology}" assumes "finite B" assumes "\x. x \ B \ set_integrable M A (f x)" shows "set_lebesgue_integral M A (\y. \x\B. f x y) = (\x\B. set_lebesgue_integral M A (f x))" using assms apply (induction rule: finite_induct) apply simp apply simp apply (subst set_integral_add) apply (auto intro!: set_integrable_sum) done lemma set_nn_integral_cong: assumes "M = M'" "A = B" "\x. x \ space M \ A \ f x = g x" shows "set_nn_integral M A f = set_nn_integral M' B g" using assms unfolding assms(1,2) by (intro nn_integral_cong) (auto simp: indicator_def) lemma set_nn_integral_mono: assumes "\x. x \ space M \ A \ f x \ g x" shows "set_nn_integral M A f \ set_nn_integral M A g" using assms by (intro nn_integral_mono) (auto simp: indicator_def) lemma fixes f :: "real \ real" assumes "a \ b" assumes deriv: "\x. a \ x \ x \ b \ (F has_field_derivative f x) (at x within {a..b})" assumes cont: "continuous_on {a..b} f" shows has_bochner_integral_FTC_Icc_real: "has_bochner_integral lborel (\x. f x * indicator {a .. b} x) (F b - F a)" (is ?has) and integral_FTC_Icc_real: "(\x. f x * indicator {a .. b} x \lborel) = F b - F a" (is ?eq) proof - have 1: "\x. a \ x \ x \ b \ (F has_vector_derivative f x) (at x within {a .. b})" unfolding has_real_derivative_iff_has_vector_derivative[symmetric] using deriv by auto show ?has ?eq using has_bochner_integral_FTC_Icc[OF \a \ b\ 1 cont] integral_FTC_Icc[OF \a \ b\ 1 cont] by (auto simp: mult.commute) qed lemma integral_by_parts_integrable: fixes f g F G::"real \ real" assumes "a \ b" assumes cont_f[intro]: "continuous_on {a..b} f" assumes cont_g[intro]: "continuous_on {a..b} g" assumes [intro]: "\x. x \ {a..b} \ (F has_field_derivative f x) (at x within {a..b})" assumes [intro]: "\x. x \ {a..b} \ (G has_field_derivative g x) (at x within {a..b})" shows "integrable lborel (\x.((F x) * (g x) + (f x) * (G x)) * indicator {a .. b} x)" proof - have "integrable lborel (\x. indicator {a..b} x *\<^sub>R ((F x) * (g x) + (f x) * (G x)))" by (intro borel_integrable_compact continuous_intros assms) (auto intro!: DERIV_continuous_on assms) thus ?thesis by (simp add: mult_ac) qed lemma integral_by_parts: fixes f g F G::"real \ real" assumes [arith]: "a \ b" assumes cont_f[intro]: "continuous_on {a..b} f" assumes cont_g[intro]: "continuous_on {a..b} g" assumes [intro]: "\x. x \ {a..b} \ (F has_field_derivative f x) (at x within {a..b})" assumes [intro]: "\x. x \ {a..b} \ (G has_field_derivative g x) (at x within {a..b})" shows "(\x. (F x * g x) * indicator {a .. b} x \lborel) = F b * G b - F a * G a - \x. (f x * G x) * indicator {a .. b} x \lborel" proof- have 0: "(\x. (F x * g x + f x * G x) * indicator {a .. b} x \lborel) = F b * G b - F a * G a" by (rule integral_FTC_Icc_real, auto intro!: derivative_eq_intros continuous_intros) (auto intro!: assms DERIV_continuous_on) have [continuous_intros]: "continuous_on {a..b} F" by (rule DERIV_continuous_on assms)+ have [continuous_intros]: "continuous_on {a..b} G" by (rule DERIV_continuous_on assms)+ have "(\x. indicator {a..b} x *\<^sub>R (F x * g x + f x * G x) \lborel) = (\x. indicator {a..b} x *\<^sub>R(F x * g x) \lborel) + \x. indicator {a..b} x *\<^sub>R (f x * G x) \lborel" apply (subst Bochner_Integration.integral_add[symmetric]) apply (rule borel_integrable_compact; force intro!: continuous_intros assms) apply (rule borel_integrable_compact; force intro!: continuous_intros assms) apply (simp add: algebra_simps) done thus ?thesis using 0 by (simp add: algebra_simps) qed lemma interval_lebesgue_integral_by_parts: assumes "a \ b" assumes cont_f[intro]: "continuous_on {a..b} f" assumes cont_g[intro]: "continuous_on {a..b} g" assumes [intro]: "\x. x \ {a..b} \ (F has_field_derivative f x) (at x within {a..b})" assumes [intro]: "\x. x \ {a..b} \ (G has_field_derivative g x) (at x within {a..b})" shows "(LBINT x=a..b. F x * g x) = F b * G b - F a * G a - (LBINT x=a..b. f x * G x)" using integral_by_parts[of a b f g F G] assms by (simp add: interval_integral_Icc set_lebesgue_integral_def mult_ac) lemma interval_lebesgue_integral_by_parts_01: assumes cont_f[intro]: "continuous_on {0..1} f" assumes cont_g[intro]: "continuous_on {0..1} g" assumes [intro]: "\x. x \ {0..1} \ (F has_field_derivative f x) (at x within {0..1})" assumes [intro]: "\x. x \ {0..1} \ (G has_field_derivative g x) (at x within {0..1})" shows "(LBINT x=0..1. F x * g x) = F 1 * G 1 - F 0 * G 0 - (LBINT x=0..1. f x * G x)" using interval_lebesgue_integral_by_parts[of 0 1 f g F G] assms by (simp add: zero_ereal_def one_ereal_def) lemma continuous_on_imp_set_integrable_cbox: fixes h :: "'a :: euclidean_space \ real" assumes "continuous_on (cbox a b) h" shows "set_integrable lborel (cbox a b) h" proof - from assms have "h absolutely_integrable_on cbox a b" by (rule absolutely_integrable_continuous) moreover have "(\x. indicat_real (cbox a b) x *\<^sub>R h x) \ borel_measurable borel" by (rule borel_measurable_continuous_on_indicator) (use assms in auto) ultimately show ?thesis unfolding set_integrable_def using assms by (subst (asm) integrable_completion) auto qed subsection \Shifted Legendre polynomials\ text \ The first ingredient we need to show Apéry's theorem is the \<^emph>\shifted Legendre polynomials\ \[P_n(X) = \frac{1}{n!} \frac{\partial^n}{\partial X^n} (X^n(1-X)^n)\] and the auxiliary polynomials \[Q_{n,k}(X) = \frac{\partial^k}{\partial X^k} (X^n(1-X)^n)\ .\] Note that $P_n$ is in fact an \emph{integer} polynomial. Only some very basic properties of these will be proven, since that is all we will need. \ context fixes n :: nat begin definition gen_shleg_poly :: "nat \ int poly" where "gen_shleg_poly k = (pderiv ^^ k) ([:0, 1, -1:] ^ n)" definition shleg_poly where "shleg_poly = gen_shleg_poly n div [:fact n:]" text \ We can easily prove the following more explicit formula for $Q_{n,k}$: \[Q_{n,k}(X) = \sum_{i=0}^k (-1)^{k-1} {k \choose i} n^{\underline{i}}\, n^{\underline{k-i}}\, X^{n-i}\, (1-X)^{n-k+i}\] \ lemma gen_shleg_poly_altdef: assumes "k \ n" shows "gen_shleg_poly k = (\i\k. smult ((-1)^(k-i) * of_nat (k choose i) * pochhammer (n-i+1) i * pochhammer (n-k+i+1) (k-i)) ([:0, 1:] ^ (n - i) * [:1, -1:] ^ (n-k+i)))" proof - have *: "(pderiv ^^ i) (x \\<^sub>p [:1, -1:]) = smult ((-1) ^ i) ((pderiv ^^ i) x \\<^sub>p [:1, -1:])" for i and x :: "int poly" by (induction i arbitrary: x) (auto simp: pderiv_smult pderiv_pcompose funpow_Suc_right pderiv_pCons higher_pderiv_minus simp del: funpow.simps(2)) have "gen_shleg_poly k = (pderiv ^^ k) ([:0, 1, -1:] ^ n)" by (simp add: gen_shleg_poly_def) also have "[:0, 1, -1::int:] = [:0, 1:] * [:1, -1:]" by simp also have "\ ^ n = [:0, 1:] ^ n * [:1, -1:] ^ n" by (simp flip: power_mult_distrib) also have "(pderiv ^^ k) \ = (\i\k. smult (of_nat (k choose i)) ((pderiv ^^ i) ([:0, 1:] ^ n) * (pderiv ^^ (k - i)) ([:1, -1:] ^ n)))" by (simp add: higher_pderiv_mult) also have "\ = (\i\k. smult (of_nat (k choose i)) ((pderiv ^^ i) (monom 1 n) * (pderiv ^^ (k - i)) (monom 1 n \\<^sub>p [:1, -1:])))" by (simp add: monom_altdef pcompose_pCons pcompose_power_left) also have "\ = (\i\k. smult ((-1) ^ (k - i) * of_nat (k choose i)) ((pderiv ^^ i) (monom 1 n) * ((pderiv ^^ (k - i)) (monom 1 n) \\<^sub>p [:1, -1:])))" by (simp add: * mult_ac) also have "\ = (\i\k. smult ((-1) ^ (k - i) * of_nat (k choose i)) (monom (pochhammer (n - i + 1) i) (n - i) * monom (pochhammer (n - k + i + 1) (k - i)) (n - k + i) \\<^sub>p [:1, -1:]))" using assms by (simp add: higher_pderiv_monom) also have "\ = (\i\k. smult ((-1) ^ (k - i) * of_nat (k choose i) * pochhammer (n - i + 1) i * pochhammer (n - k + i + 1) (k - i)) ([:0, 1:] ^ (n - i) * [:1, -1:] ^ (n - k + i)))" by (simp add: monom_altdef algebra_simps pcompose_smult pcompose_power_left pcompose_pCons) finally show ?thesis . qed lemma degree_gen_shleg_poly [simp]: "degree (gen_shleg_poly k) = 2 * n - k" by (simp add: gen_shleg_poly_def degree_higher_pderiv degree_power_eq) lemma gen_shleg_poly_n: "gen_shleg_poly n = smult (fact n) shleg_poly" proof - obtain r where r: "gen_shleg_poly n = [:fact n:] * r" unfolding gen_shleg_poly_def using fact_dvd_higher_pderiv[of n "[:0,1,-1:]^n"] by blast have "smult (fact n) shleg_poly = smult (fact n) (gen_shleg_poly n div [:fact n:])" by (simp add: shleg_poly_def) also note r also have "[:fact n:] * r div [:fact n:] = r" by (rule nonzero_mult_div_cancel_left) auto finally show ?thesis by (simp add: r) qed lemma degree_shleg_poly [simp]: "degree shleg_poly = n" using degree_gen_shleg_poly[of n] by (simp add: gen_shleg_poly_n) lemma pderiv_gen_shleg_poly [simp]: "pderiv (gen_shleg_poly k) = gen_shleg_poly (Suc k)" by (simp add: gen_shleg_poly_def) text \ The following functions are the interpretation of the shifted Legendre polynomials and the auxiliary polynomials as a function from reals to reals. \ definition Gen_Shleg :: "nat \ real \ real" where "Gen_Shleg k x = poly (of_int_poly (gen_shleg_poly k)) x" definition Shleg :: "real \ real" where "Shleg = poly (of_int_poly shleg_poly)" lemma Gen_Shleg_altdef: assumes "k \ n" shows "Gen_Shleg k x = (\i\k. (-1)^(k-i) * of_nat (k choose i) * of_int (pochhammer (n-i+1) i * pochhammer (n-k+i+1) (k-i)) * x ^ (n - i) * (1 - x) ^ (n-k+i))" using assms by (simp add: Gen_Shleg_def gen_shleg_poly_altdef poly_sum mult_ac) lemma Gen_Shleg_0 [simp]: "k < n \ Gen_Shleg k 0 = 0" by (simp add: Gen_Shleg_altdef zero_power) lemma Gen_Shleg_1 [simp]: "k < n \ Gen_Shleg k 1 = 0" by (simp add: Gen_Shleg_altdef zero_power) lemma Gen_Shleg_n_0 [simp]: "Gen_Shleg n 0 = fact n" proof - have "Gen_Shleg n 0 = (\i\n. (-1) ^ (n-i) * real (n choose i) * (real (pochhammer (Suc (n-i)) i) * real (pochhammer (Suc i) (n-i))) * 0 ^ (n - i))" by (simp add: Gen_Shleg_altdef) also have "\ = (\i\{n}. (-1) ^ (n-i) * real (n choose i) * (real (pochhammer (Suc (n-i)) i) * real (pochhammer (Suc i) (n-i))) * 0 ^ (n - i))" by (intro sum.mono_neutral_right) auto also have "\ = fact n" by (simp add: pochhammer_fact flip: pochhammer_of_nat) finally show ?thesis . qed lemma Gen_Shleg_n_1 [simp]: "Gen_Shleg n 1 = (-1) ^ n * fact n" proof - have "Gen_Shleg n 1 = (\i\n. (-1) ^ (n-i) * real (n choose i) * (real (pochhammer (Suc (n-i)) i) * real (pochhammer (Suc i) (n-i))) * 0 ^ i)" by (simp add: Gen_Shleg_altdef) also have "\ = (\i\{0}. (-1) ^ (n-i) * real (n choose i) * (real (pochhammer (Suc (n-i)) i) * real (pochhammer (Suc i) (n-i))) * 0 ^ i)" by (intro sum.mono_neutral_right) auto also have "\ = (-1) ^ n * fact n" by (simp add: pochhammer_fact flip: pochhammer_of_nat) finally show ?thesis . qed lemma Shleg_altdef: "Shleg x = Gen_Shleg n x / fact n" by (simp add: Shleg_def Gen_Shleg_def gen_shleg_poly_n) lemma Shleg_0 [simp]: "Shleg 0 = 1" and Shleg_1 [simp]: "Shleg 1 = (-1) ^ n" by (simp_all add: Shleg_altdef) lemma Gen_Shleg_0_left: "Gen_Shleg 0 x = x ^ n * (1 - x) ^ n" by (simp add: Gen_Shleg_def gen_shleg_poly_def power_mult_distrib) lemma has_field_derivative_Gen_Shleg: "(Gen_Shleg k has_field_derivative Gen_Shleg (Suc k) x) (at x)" proof - note [derivative_intros] = poly_DERIV show ?thesis unfolding Gen_Shleg_def by (rule derivative_eq_intros) auto qed lemma continuous_on_Gen_Shleg: "continuous_on A (Gen_Shleg k)" by (auto simp: Gen_Shleg_def intro!: continuous_intros) lemma continuous_on_Gen_Shleg' [continuous_intros]: "continuous_on A f \ continuous_on A (\x. Gen_Shleg k (f x))" by (rule continuous_on_compose2[OF continuous_on_Gen_Shleg[of UNIV]]) auto lemma continuous_on_Shleg: "continuous_on A Shleg" by (auto simp: Shleg_def intro!: continuous_intros) lemma continuous_on_Shleg' [continuous_intros]: "continuous_on A f \ continuous_on A (\x. Shleg (f x))" by (rule continuous_on_compose2[OF continuous_on_Shleg[of UNIV]]) auto lemma measurable_Gen_Shleg [measurable]: "Gen_Shleg n \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_on_Gen_Shleg) lemma measurable_Shleg [measurable]: "Shleg \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_on_Shleg) end subsection \Auxiliary facts about the \\\ function\ lemma Re_zeta_ge_1: assumes "x > 1" shows "Re (zeta (of_real x)) \ 1" proof - have *: "(\n. real (Suc n) powr -x) sums Re (zeta (complex_of_real x))" using sums_Re[OF sums_zeta[of "of_real x"]] assms by (simp add: powr_Reals_eq) show "Re (zeta (of_real x)) \ 1" proof (rule sums_le[OF _ _ *]) show "(\n. if n = 0 then 1 else 0) sums 1" by (rule sums_single) qed auto qed lemma sums_zeta_of_nat_offset: fixes r :: nat assumes n: "n > 1" shows "(\k. 1 / (r + k + 1) ^ n) sums (zeta (of_nat n) - (\k=1..r. 1 / k ^ n))" proof - have "(\k. 1 / (k + 1) ^ n) sums zeta (of_nat n)" using sums_zeta[of "of_nat n"] n by (simp add: powr_minus field_simps flip: of_nat_Suc) from sums_split_initial_segment[OF this, of r] have "(\k. 1 / (r + k + 1) ^ n) sums (zeta (of_nat n) - (\kkk=1..r. 1 / k ^ n)" by (intro sum.reindex_bij_witness[of _ "\k. k - 1" Suc]) auto finally show ?thesis . qed lemma sums_Re_zeta_of_nat_offset: fixes r :: nat assumes n: "n > 1" shows "(\k. 1 / (r + k + 1) ^ n) sums (Re (zeta (of_nat n)) - (\k=1..r. 1 / k ^ n))" proof - have "(\k. Re (1 / (r + k + 1) ^ n)) sums (Re (zeta (of_nat n) - (\k=1..r. 1 / k ^ n)))" by (intro sums_Re sums_zeta_of_nat_offset assms) thus ?thesis by simp qed subsection \Divisor of a sum of rationals\ text \ A finite sum of rationals of the form $\frac{a_1}{b_1} + \ldots + \frac{a_n}{b_n}$ can be brought into the form $\frac{c}{d}$, where $d$ is the LCM of the $b_i$ (or some integer multiple thereof). \ lemma sum_rationals_common_divisor: fixes f g :: "'a \ int" assumes "finite A" assumes "\x. x \ A \ g x \ 0" shows "\c. (\x\A. f x / g x) = real_of_int c / (LCM x\A. g x)" using assms proof (induction rule: finite_induct) case empty thus ?case by auto next case (insert x A) define d where "d = (LCM x\A. g x)" from insert have [simp]: "d \ 0" by (auto simp: d_def Lcm_0_iff) from insert have [simp]: "g x \ 0" by auto from insert obtain c where c: "(\x\A. f x / g x) = real_of_int c / real_of_int d" by (auto simp: d_def) define e1 where "e1 = lcm d (g x) div d" define e2 where "e2 = lcm d (g x) div g x" have "(\y\insert x A. f y / g y) = c / d + f x / g x" using insert c by simp also have "c / d = (c * e1) / lcm d (g x)" by (simp add: e1_def real_of_int_div) also have "f x / g x = (f x * e2) / lcm d (g x)" by (simp add: e2_def real_of_int_div) also have "(c * e1) / lcm d (g x) + \ = (c * e1 + f x * e2) / (LCM x\insert x A. g x)" using insert by (simp add: add_divide_distrib lcm.commute d_def) finally show ?case .. qed lemma sum_rationals_common_divisor': fixes f g :: "'a \ int" assumes "finite A" assumes "\x. x \ A \ g x \ 0" and "(\x. x \ A \ g x dvd d)" and "d \ 0" shows "\c. (\x\A. f x / g x) = real_of_int c / real_of_int d" proof - define d' where "d' = (LCM x\A. g x)" have "d' dvd d" unfolding d'_def using assms(3) by (auto simp: Lcm_dvd_iff) then obtain e where e: "d = d' * e" by blast have "\c. (\x\A. f x / g x) = real_of_int c / (LCM x\A. g x)" by (rule sum_rationals_common_divisor) fact+ then obtain c where c: "(\x\A. f x / g x) = real_of_int c / real_of_int d'" unfolding d'_def .. also have "\ = real_of_int (c * e) / real_of_int d" using \d \ 0\ by (simp add: e) finally show ?thesis .. qed subsection \The first double integral\ text \ We shall now investigate the double integral \[I_1 := \int_0^1 \int_0^1 -\frac{\ln(xy)}{1-xy}\,x^r y^s\, \text{d}x\,\text{d}y\ .\] Since everything is non-negative for now, we can work over the extended non-negative real numbers and the issues of integrability or summability do not arise at all. \ definition beukers_nn_integral1 :: "nat \ nat \ ennreal" where "beukers_nn_integral1 r s = (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. ennreal (-ln (x*y) / (1 - x*y) * x^r * y^s) \lborel)" definition beukers_integral1 :: "nat \ nat \ real" where "beukers_integral1 r s = (\(x,y)\{0<..<1}\{0<..<1}. (-ln (x*y) / (1 - x*y) * x^r * y^s) \lborel)" lemma fixes x y z :: real assumes xyz: "x \ {0<..<1}" "y \ {0<..<1}" "z \ {0..1}" shows beukers_denom_ineq: "(1 - x * y) * z < 1" and beukers_denom_neq: "(1 - x * y) * z \ 1" proof - from xyz have *: "x * y < 1 * 1" by (intro mult_strict_mono) auto from * have "(1 - x * y) * z \ (1 - x * y) * 1" using xyz by (intro mult_left_mono) auto also have "\ < 1 * 1" using xyz by (intro mult_strict_right_mono) auto finally show "(1 - x * y) * z < 1" "(1 - x * y) * z \ 1" by simp_all qed text \ We first evaluate the improper integral \[\int_0^1 -\ln x \cdot x^e\,\text{d}x = \frac{1}{(e+1)^2}\ .\] for any $e>-1$. \ lemma integral_0_1_ln_times_powr: assumes "e > -1" shows "(LBINT x=0..1. -ln x * x powr e) = 1 / (e + 1)\<^sup>2" and "interval_lebesgue_integrable lborel 0 1 (\x. -ln x * x powr e)" proof - define f where "f = (\x. -ln x * x powr e)" define F where "F = (\x. x powr (e + 1) * (1 - (e + 1) * ln x) / (e + 1) ^ 2)" have 0: "isCont f x" if "x \ {0<..<1}" for x using that by (auto intro!: continuous_intros simp: f_def) have 1: "(F has_real_derivative f x) (at x)" if "x \ {0<..<1}" for x proof - show "(F has_real_derivative f x) (at x)" unfolding F_def f_def using that assms apply (insert that assms) apply (rule derivative_eq_intros refl | simp)+ apply (simp add: divide_simps) apply (simp add: power2_eq_square algebra_simps powr_add power_numeral_reduce) done qed have 2: "AE x in lborel. ereal 0 < ereal x \ ereal x < ereal 1 \ 0 \ f x" by (intro AE_I2) (auto simp: f_def mult_nonpos_nonneg) have 3: "((F \ real_of_ereal) \ 0) (at_right (ereal 0))" unfolding ereal_tendsto_simps F_def using assms by real_asymp have 4: "((F \ real_of_ereal) \ F 1) (at_left (ereal 1))" unfolding ereal_tendsto_simps F_def using assms by real_asymp (simp add: field_simps) have "(LBINT x=ereal 0..ereal 1. f x) = F 1 - 0" by (rule interval_integral_FTC_nonneg[where F = F]) (use 0 1 2 3 4 in auto) thus "(LBINT x=0..1. -ln x * x powr e) = 1 / (e + 1)\<^sup>2" by (simp add: F_def zero_ereal_def one_ereal_def f_def) have "set_integrable lborel (einterval (ereal 0) (ereal 1)) f" by (rule interval_integral_FTC_nonneg) (use 0 1 2 3 4 in auto) thus "interval_lebesgue_integrable lborel 0 1 f" by (simp add: interval_lebesgue_integrable_def einterval_def) qed lemma interval_lebesgue_integral_lborel_01_cong: assumes "\x. x \ {0<..<1} \ f x = g x" shows "interval_lebesgue_integral lborel 0 1 f = interval_lebesgue_integral lborel 0 1 g" using assms by (subst (1 2) interval_integral_Ioo) (auto intro!: set_lebesgue_integral_cong assms) lemma nn_integral_0_1_ln_times_powr: assumes "e > -1" shows "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * y powr e) \lborel) = ennreal (1 / (e + 1)\<^sup>2)" proof - have *: "(LBINT x=0..1. - ln x * x powr e = 1 / (e + 1)\<^sup>2)" "interval_lebesgue_integrable lborel 0 1 (\x. - ln x * x powr e)" using integral_0_1_ln_times_powr[OF assms] by simp_all have eq: "(\y. (if 0 < y \ y < 1 then 1 else 0) * ln y * y powr e) = (\y. if 0 < y \ y < 1 then ln y * y powr e else 0)" by auto have "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * y powr e) \lborel) = (\\<^sup>+y. ennreal (-ln y * y powr e * indicator {0<..<1} y) \lborel)" by (intro nn_integral_cong) (auto simp: indicator_def) also have "\ = ennreal (1 / (e + 1)\<^sup>2)" using * eq by (subst nn_integral_eq_integral) (auto intro!: AE_I2 simp: indicator_def interval_lebesgue_integrable_def set_integrable_def one_ereal_def zero_ereal_def interval_integral_Ioo mult_ac mult_nonpos_nonneg set_lebesgue_integral_def) finally show ?thesis . qed lemma nn_integral_0_1_ln_times_power: "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * y ^ n) \lborel) = ennreal (1 / (n + 1)\<^sup>2)" proof - have "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * y ^ n) \lborel) = (\\<^sup>+y\{0<..<1}. ennreal (-ln y * y powr real n) \lborel)" by (intro set_nn_integral_cong) (auto simp: powr_realpow) also have "\ = ennreal (1 / (n + 1)^2)" by (subst nn_integral_0_1_ln_times_powr) auto finally show ?thesis by simp qed text \ Next, we also evaluate the more trivial integral \[\int_0^1 x^n\ \text{d}x\ .\] \ lemma nn_integral_0_1_power: "(\\<^sup>+y\{0<..<1}. ennreal (y ^ n) \lborel) = ennreal (1 / (n + 1))" proof - have *: "((\a. a ^ (n + 1) / real (n + 1)) has_real_derivative x ^ n) (at x)" for x by (rule derivative_eq_intros refl | simp)+ have "(\\<^sup>+y\{0<..<1}. ennreal (y ^ n) \lborel) = (\\<^sup>+y\{0..1}. ennreal (y ^ n) \lborel)" by (intro nn_integral_cong_AE AE_I[of _ _ "{0,1}"]) (auto simp: indicator_def emeasure_lborel_countable) also have "\ = ennreal (1 ^ (n + 1) / (n + 1) - 0 ^ (n + 1) / (n + 1))" using * by (intro nn_integral_FTC_Icc) auto also have "\ = ennreal (1 / (n + 1))" by simp finally show ?thesis by simp qed text \ $I_1$ can alternatively be written as the triple integral \[\int_0^1\int_0^1\int_0^1 \frac{x^r y^s}{1-(1-xy)w}\ \text{d}x\,\text{d}y\,\text{d}w\ .\] \ lemma beukers_nn_integral1_altdef: "beukers_nn_integral1 r s = (\\<^sup>+(w,x,y)\{0<..<1}\{0<..<1}\{0<..<1}. ennreal (1 / (1-(1-x*y)*w) * x^r * y^s) \lborel)" proof - have "(\\<^sup>+(w,x,y)\{0<..<1}\{0<..<1}\{0<..<1}. ennreal (1 / (1-(1-x*y)*w) * x^r * y^s) \lborel) = (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. (\\<^sup>+w\{0<..<1}. ennreal (1 / (1-(1-x*y)*w) * x^r * y^s) \lborel) \lborel)" by (subst lborel_prod [symmetric], subst lborel_pair.nn_integral_snd [symmetric]) (auto simp: case_prod_unfold indicator_def simp flip: lborel_prod intro!: nn_integral_cong) also have "\ = (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. ennreal (-ln (x*y)/(1-x*y) * x^r * y^s) \lborel)" proof (intro nn_integral_cong, clarify) fix x y :: real have "(\\<^sup>+w\{0<..<1}. ennreal (1/(1-(1-x*y)*w)*x^r*y^s) \lborel) = ennreal (-ln (x*y)*x^r*y^s/(1-x*y))" if xy: "(x, y) \ {0<..<1} \ {0<..<1}" proof - from xy have "x * y < 1" using mult_strict_mono[of x 1 y 1] by simp have deriv: "((\w. -ln (1-(1-x*y)*w) / (1-x*y)) has_real_derivative 1/(1-(1-x*y)*w)) (at w)" if w: "w \ {0..1}" for w by (insert xy w \x*y<1\ beukers_denom_ineq[of x y w]) (rule derivative_eq_intros refl | simp add: divide_simps)+ have "(\\<^sup>+w\{0<..<1}. ennreal (1/(1-(1-x*y)*w)*x^r*y^s) \lborel) = ennreal (x^r*y^s) * (\\<^sup>+w\{0<..<1}. ennreal (1/(1-(1-x*y)*w)) \lborel)" using xy by (subst nn_integral_cmult [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def simp flip: ennreal_mult') also have "(\\<^sup>+w\{0<..<1}. ennreal (1/(1-(1-x*y)*w)) \lborel) = (\\<^sup>+w\{0..1}. ennreal (1/(1-(1-x*y)*w)) \lborel)" by (intro nn_integral_cong_AE AE_I[of _ _ "{0,1}"]) (auto simp: emeasure_lborel_countable indicator_def) also have "(\\<^sup>+w\{0..1}. ennreal (1/(1-(1-x*y)*w)) \lborel) = ennreal (-ln (1-(1-x*y)*1)/(1-x*y) - (-ln (1-(1-x*y)*0)/(1-x*y)))" using xy deriv less_imp_le[OF beukers_denom_ineq[of x y]] by (intro nn_integral_FTC_Icc) auto finally show ?thesis using xy by (simp flip: ennreal_mult' ennreal_mult'' add: mult_ac) qed thus "(\\<^sup>+w\{0<..<1}. ennreal (1/(1-(1-x*y)*w)*x^r*y^s) \lborel) * indicator ({0<..<1}\{0<..<1}) (x, y) = ennreal (-ln (x*y)/(1-x*y)*x^r*y^s) * indicator ({0<..<1}\{0<..<1}) (x, y)" by (auto simp: indicator_def) qed also have "\ = beukers_nn_integral1 r s" by (simp add: beukers_nn_integral1_def) finally show ?thesis .. qed context fixes r s :: nat and I1 I2' :: real and I2 :: ennreal and D :: "(real \ real \ real) set" assumes rs: "s \ r" defines "D \ {0<..<1}\{0<..<1}\{0<..<1}" begin text \ By unfolding the geometric series, pulling the summation out and evaluating the integrals, we find that \[I_1 = \sum_{k=0}^\infty \frac{1}{(k+r+1)^2(k+s+1)} + \frac{1}{(k+r+1)(k+s+1)^2}\ .\] \ lemma beukers_nn_integral1_series: "beukers_nn_integral1 r s = (\k. ennreal (1/((k+r+1)^2*(k+s+1)) + 1/((k+r+1)*(k+s+1)^2)))" proof - have "beukers_nn_integral1 r s = (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. (\k. ennreal (-ln (x*y) * x^(k+r) * y^(k+s))) \lborel)" unfolding beukers_nn_integral1_def proof (intro set_nn_integral_cong refl, clarify) fix x y :: real assume xy: "x \ {0<..<1}" "y \ {0<..<1}" from xy have "x * y < 1" using mult_strict_mono[of x 1 y 1] by simp have "(\k. ennreal (-ln (x*y) * x^(k+r) * y^(k+s))) = ennreal (-ln (x*y) * x^r * y^s) * (\k. ennreal ((x*y)^k))" using xy by (subst ennreal_suminf_cmult [symmetric], subst ennreal_mult'' [symmetric]) (auto simp: power_add mult_ac power_mult_distrib) also have "(\k. ennreal ((x*y)^k)) = ennreal (1 / (1 - x*y))" using geometric_sums[of "x*y"] \x * y < 1\ xy by (intro suminf_ennreal_eq) auto also have "ennreal (-ln (x*y) * x^r * y^s) * \ = ennreal (-ln (x*y) / (1 - x*y) * x^r * y^s)" using \x * y < 1\ by (subst ennreal_mult'' [symmetric]) auto finally show "ennreal (-ln (x*y) / (1 - x*y) * x^r * y^s) = (\k. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)))" .. qed also have "\ = (\k. (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. (ennreal (-ln (x*y) * x^(k+r) * y^(k+s))) \lborel))" unfolding case_prod_unfold by (subst nn_integral_suminf [symmetric]) (auto simp flip: borel_prod) also have "\ = (\k. ennreal (1 / ((k+r+1)^2*(k+s+1)) + 1 / ((k+r+1)*(k+s+1)^2)))" proof (rule suminf_cong) fix k :: nat define F where "F = (\x y::real. x + y)" have "(\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)) \lborel) = (\\<^sup>+x\{0<..<1}. (\\<^sup>+y\{0<..<1}. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)) \lborel) \lborel)" unfolding case_prod_unfold lborel_prod [symmetric] by (subst lborel.nn_integral_fst [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def) also have "\ = (\\<^sup>+x\{0<..<1}. ennreal (-ln x * x^(k+r) / (k+s+1) + x^(k+r)/(k+s+1)^2) \lborel)" proof (intro set_nn_integral_cong refl, clarify) fix x :: real assume x: "x \ {0<..<1}" have "(\\<^sup>+y\{0<..<1}. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)) \lborel) = (\\<^sup>+y\{0<..<1}. (ennreal (-ln x * x^(k+r) * y^(k+s)) + ennreal (-ln y * x^(k+r) * y^(k+s))) \lborel)" by (intro set_nn_integral_cong) (use x in \auto simp: ln_mult ring_distribs mult_nonpos_nonneg simp flip: ennreal_plus\) also have "\ = (\\<^sup>+y\{0<..<1}. ennreal (-ln x * x^(k+r) * y^(k+s)) \lborel) + (\\<^sup>+y\{0<..<1}. ennreal (-ln y * x^(k+r) * y^(k+s)) \lborel)" by (subst nn_integral_add [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def) also have "(\\<^sup>+y\{0<..<1}. ennreal (-ln x * x^(k+r) * y^(k+s)) \lborel) = ennreal (-ln x * x^(k+r)) * (\\<^sup>+y\{0<..<1}. ennreal (y^(k+s)) \lborel)" by (subst nn_integral_cmult [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def simp flip: ennreal_mult'') also have "(\\<^sup>+y\{0<..<1}. ennreal (y^(k+s)) \lborel) = ennreal (1/(k+s+1))" by (subst nn_integral_0_1_power) simp also have "ennreal (-ln x * x^(k+r)) * \ = ennreal (-ln x * x^(k+r) / (k+s+1))" by (subst ennreal_mult'' [symmetric]) auto also have "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * x^(k+r) * y^(k+s)) \lborel) = ennreal (x^(k+r)) * (\\<^sup>+y\{0<..<1}. ennreal (-ln y * y^(k+s)) \lborel)" by (subst nn_integral_cmult [symmetric]) (use x in \auto intro!: nn_integral_cong simp: indicator_def mult_ac simp flip: ennreal_mult'\) also have "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * y^(k+s)) \lborel) = ennreal (1 / (k + s + 1)\<^sup>2)" by (subst nn_integral_0_1_ln_times_power) simp also have "ennreal (x ^ (k + r)) * \ = ennreal (x ^ (k + r) / (k + s + 1) ^ 2)" by (subst ennreal_mult'' [symmetric]) auto also have "ennreal (- ln x * x ^ (k + r) / (k + s + 1)) + \ = ennreal (-ln x * x^(k+r) / (k+s+1) + x^(k+r)/(k+s+1)^2)" using x by (subst ennreal_plus) (auto simp: mult_nonpos_nonneg divide_nonpos_nonneg) finally show "(\\<^sup>+y\{0<..<1}. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)) \lborel) = ennreal (-ln x * x^(k+r) / (k+s+1) + x^(k+r)/(k+s+1)^2)" . qed also have "\ = (\\<^sup>+x\{0<..<1}. (ennreal (-ln x * x^(k+r) / (k+s+1)) + ennreal (x^(k+r)/(k+s+1)^2)) \lborel)" by (intro set_nn_integral_cong refl, subst ennreal_plus) (auto simp: mult_nonpos_nonneg divide_nonpos_nonneg) also have "\ = (\\<^sup>+x\{0<..<1}. ennreal (-ln x * x^(k+r) / (k+s+1)) \lborel) + (\\<^sup>+x\{0<..<1}. ennreal (x^(k+r)/(k+s+1)^2) \lborel)" by (subst nn_integral_add [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def) also have "(\\<^sup>+x\{0<..<1}. ennreal (-ln x * x^(k+r) / (k+s+1)) \lborel) = ennreal (1 / (k+s+1)) * (\\<^sup>+x\{0<..<1}. ennreal (-ln x * x^(k+r)) \lborel)" by (subst nn_integral_cmult [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def simp flip: ennreal_mult') also have "\ = ennreal (1 / ((k+s+1) * (k+r+1)^2))" by (subst nn_integral_0_1_ln_times_power, subst ennreal_mult [symmetric]) (auto simp: algebra_simps) also have "(\\<^sup>+x\{0<..<1}. ennreal (x^(k+r)/(k+s+1)^2) \lborel) = ennreal (1/(k+s+1)^2) * (\\<^sup>+x\{0<..<1}. ennreal (x^(k+r)) \lborel)" by (subst nn_integral_cmult [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def simp flip: ennreal_mult') also have "\ = ennreal (1/((k+r+1)*(k+s+1)^2))" by (subst nn_integral_0_1_power, subst ennreal_mult [symmetric]) (auto simp: algebra_simps) also have "ennreal (1 / ((k+s+1) * (k+r+1)^2)) + \ = ennreal (1/((k+r+1)^2*(k+s+1)) + 1/((k+r+1)*(k+s+1)^2))" by (subst ennreal_plus [symmetric]) (auto simp: algebra_simps) finally show "(\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)) \lborel) = \" . qed finally show ?thesis . qed text \ Remembering that $\zeta(3) = \sum k^{-3}$, it is easy to see that if $r = s$, this sum is simply \[2\left(\zeta(3) - \sum_{k=1}^r \frac{1}{k^3}\right)\ .\] \ lemma beukers_nn_integral1_same: assumes "r = s" shows "beukers_nn_integral1 r s = ennreal (2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3)))" and "2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3)) \ 0" proof - from assms have [simp]: "s = r" by simp have *: "Suc 2 = 3" by simp have "beukers_nn_integral1 r s = (\k. ennreal (2 / (r + k + 1) ^ 3))" unfolding beukers_nn_integral1_series by (simp only: assms power_Suc [symmetric] mult.commute[of "x ^ 2" for x] * times_divide_eq_right mult_1_right add_ac flip: mult_2) also have **: "(\k. 2 / (r + k + 1) ^ 3) sums (2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3)))" using sums_mult[OF sums_Re_zeta_of_nat_offset[of 3], of 2] by simp hence "(\k. ennreal (2 / (r + k + 1) ^ 3)) = ennreal \" by (intro suminf_ennreal_eq) auto finally show "beukers_nn_integral1 r s = ennreal (2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3)))" . show "2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3)) \ 0" by (rule sums_le[OF _ sums_zero **]) auto qed lemma beukers_integral1_same: assumes "r = s" shows "beukers_integral1 r s = 2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3))" proof - have "ln (a * b) * a ^ r * b ^ s / (1 - a * b) \ 0" if "a \ {0<..<1}" "b \ {0<..<1}" for a b :: real using that mult_strict_mono[of a 1 b 1] by (intro mult_nonpos_nonneg divide_nonpos_nonneg) auto thus ?thesis using beukers_nn_integral1_same[OF assms] unfolding beukers_nn_integral1_def beukers_integral1_def by (intro set_integral_eq_nn_integral AE_I2) (auto simp flip: lborel_prod simp: case_prod_unfold set_borel_measurable_def intro: divide_nonpos_nonneg mult_nonpos_nonneg) qed text \ In contrast, for \r > s\, we find that \[I_1 = \frac{1}{r-s} \sum_{k=s+1}^r \frac{1}{k^2}\ .\] \ lemma beukers_nn_integral1_different: assumes "r > s" shows "beukers_nn_integral1 r s = ennreal ((\k\{s<..r}. 1 / k ^ 2) / (r - s))" proof - have "(\k. 1 / (r - s) * (1 / (s + k + 1) ^ 2 - 1 / (r + k + 1) ^ 2)) sums (1 / (r - s) * ((Re (zeta (of_nat 2)) - (\k=1..s. 1/k^2)) - (Re (zeta (of_nat 2)) - (\k=1..r. 1/k^2))))" (is "_ sums ?S") by (intro sums_mult sums_diff sums_Re_zeta_of_nat_offset) auto also have "?S = ((\k=1..r. 1/k^2) - (\k=1..s. 1/k^2)) / (r - s)" by (simp add: algebra_simps diff_divide_distrib) also have "(\k=1..r. 1/k^2) - (\k=1..s. 1/k^2) = (\k\{1..r}-{1..s}. 1/k^2)" using assms by (subst Groups_Big.sum_diff) auto also have "{1..r} - {1..s} = {s<..r}" by auto also have "(\k. 1 / (r - s) * (1 / (s + k + 1) ^ 2 - 1 / (r + k + 1) ^ 2)) = (\k. 1 / ((k+r+1) * (k+s+1)^2) + 1 / ((k+r+1)^2 * (k+s+1)))" proof (intro ext, goal_cases) case (1 k) define x where "x = real (k + r + 1)" define y where "y = real (k + s + 1)" have [simp]: "x \ 0" "y \ 0" by (auto simp: x_def y_def) have "(x\<^sup>2 * y + x * y\<^sup>2) * (real r - real s) = x * y * (x\<^sup>2 - y\<^sup>2)" by (simp add: algebra_simps power2_eq_square x_def y_def) hence "1 / (x*y^2) + 1 / (x^2*y) = 1 / (r - s) * (1 / y^2 - 1 / x^2)" using assms by (simp add: divide_simps of_nat_diff) thus ?case by (simp add: x_def y_def algebra_simps) qed finally show ?thesis unfolding beukers_nn_integral1_series by (intro suminf_ennreal_eq) (auto simp: add_ac) qed lemma beukers_integral1_different: assumes "r > s" shows "beukers_integral1 r s = (\k\{s<..r}. 1 / k ^ 2) / (r - s)" proof - have "ln (a * b) * a ^ r * b ^ s / (1 - a * b) \ 0" if "a \ {0<..<1}" "b \ {0<..<1}" for a b :: real using that mult_strict_mono[of a 1 b 1] by (intro mult_nonpos_nonneg divide_nonpos_nonneg) auto thus ?thesis using beukers_nn_integral1_different[OF assms] unfolding beukers_nn_integral1_def beukers_integral1_def by (intro set_integral_eq_nn_integral AE_I2) (auto simp flip: lborel_prod simp: case_prod_unfold set_borel_measurable_def intro: divide_nonpos_nonneg mult_nonpos_nonneg intro!: sum_nonneg divide_nonneg_nonneg) qed end text \ It is also easy to see that if we exchange \r\ and \s\, nothing changes. \ lemma beukers_nn_integral1_swap: "beukers_nn_integral1 r s = beukers_nn_integral1 s r" unfolding beukers_nn_integral1_def lborel_prod [symmetric] by (subst lborel_pair.nn_integral_swap, simp) (intro nn_integral_cong, auto simp: indicator_def algebra_simps split: if_splits) lemma beukers_nn_integral1_finite: "beukers_nn_integral1 r s < \" using beukers_nn_integral1_different[of r s] beukers_nn_integral1_different[of s r] by (cases r s rule: linorder_cases) (simp_all add: beukers_nn_integral1_same beukers_nn_integral1_swap) lemma beukers_integral1_integrable: "set_integrable lborel ({0<..<1}\{0<..<1}) (\(x,y). (-ln (x*y) / (1 - x*y) * x^r * y^s :: real))" proof (intro set_integrableI_nonneg AE_I2; clarify?) fix x y :: real assume xy: "x \ {0<..<1}" "y \ {0<..<1}" have "0 \ ln (x * y) / (1 - x * y) * x ^ r * y ^ s" using mult_strict_mono[of x 1 y 1] by (intro mult_nonpos_nonneg divide_nonpos_nonneg) (use xy in auto) thus "0 \ - ln (x * y) / (1 - x * y) * x ^ r * y ^ s" by simp next show "\\<^sup>+x\{0<..<1}\{0<..<1}. ennreal (case x of (x, y) \ - ln (x * y) / (1 - x * y) * x ^ r * y ^ s) \lborel < \" using beukers_nn_integral1_finite by (simp add: beukers_nn_integral1_def case_prod_unfold) qed (simp_all flip: lborel_prod add: set_borel_measurable_def) lemma beukers_integral1_integrable': "set_integrable lborel ({0<..<1}\{0<..<1}\{0<..<1}) (\(z,x,y). (x^r * y^s / (1 - (1 - x*y) * z) :: real))" proof (intro set_integrableI_nonneg AE_I2; clarify?) fix x y z :: real assume xyz: "x \ {0<..<1}" "y \ {0<..<1}" "z \ {0<..<1}" show "0 \ x^r * y^s / (1 - (1 - x*y) * z)" using mult_strict_mono[of x 1 y 1] xyz beukers_denom_ineq[of x y z] by (intro mult_nonneg_nonneg divide_nonneg_nonneg) auto next show "\\<^sup>+x\{0<..<1}\{0<..<1}\{0<..<1}. ennreal (case x of (z,x,y) \ x ^ r * y ^ s / (1-(1-x*y)*z)) \lborel < \" using beukers_nn_integral1_finite by (simp add: beukers_nn_integral1_altdef case_prod_unfold) qed (simp_all flip: lborel_prod add: set_borel_measurable_def) lemma beukers_integral1_conv_nn_integral: "beukers_integral1 r s = enn2real (beukers_nn_integral1 r s)" proof - have "ln (a * b) * a ^ r * b ^ s / (1 - a * b) \ 0" if "a \ {0<..<1}" "b \ {0<..<1}" for a b :: real using mult_strict_mono[of a 1 b 1] that by (intro divide_nonpos_nonneg mult_nonpos_nonneg) auto thus ?thesis unfolding beukers_integral1_def using beukers_nn_integral1_finite[of r s] by (intro set_integral_eq_nn_integral) (auto simp: case_prod_unfold beukers_nn_integral1_def set_borel_measurable_def simp flip: borel_prod intro!: AE_I2 intro: divide_nonpos_nonneg mult_nonpos_nonneg) qed lemma beukers_integral1_swap: "beukers_integral1 r s = beukers_integral1 s r" by (simp add: beukers_integral1_conv_nn_integral beukers_nn_integral1_swap) subsection \The second double integral\ context fixes n :: nat fixes D :: "(real \ real) set" and D' :: "(real \ real \ real) set" fixes P :: "real \ real" and Q :: "nat \ real \ real" defines "D \ {0<..<1} \ {0<..<1}" and "D' \ {0<..<1} \ {0<..<1} \ {0<..<1}" defines "Q \ Gen_Shleg n" and "P \ Shleg n" begin text \ The next integral to consider is the following variant of $I_1$: \[I_2 := \int_0^1 \int_0^1 -\frac{\ln(xy)}{1-xy} P_n(x) P_n(y)\,\text{d}x\,\text{d}y\ .\] \ definition beukers_integral2 :: real where "beukers_integral2 = (\(x,y)\D. (-ln (x*y) / (1-x*y) * P x * P y) \lborel)" text \ $I_2$ is simply a sum of integrals of type $I_1$, so using our results for $I_1$, we can write $I_2$ in the form $A \zeta(3) + \frac{B}{\text{lcm}\{1\ldots n\}^3}$ where $A$ and $B$ are integers and $A > 0$: \ lemma beukers_integral2_conv_int_combination: obtains A B :: int where "A > 0" and "beukers_integral2 = of_int A * Re (zeta 3) + of_int B / of_nat (Lcm {1..n} ^ 3)" proof - let ?I1 = "(\i. (i, i)) ` {..n}" let ?I2 = "Set.filter (\(i,j). i \ j) ({..n}\{..n})" let ?I3 = "Set.filter (\(i,j). i < j) ({..n}\{..n})" let ?I4 = "Set.filter (\(i,j). i > j) ({..n}\{..n})" define p where "p = shleg_poly n" define I where "I = (SIGMA i:{..n}. {1..i})" define J where "J = (SIGMA (i,j):?I4. {j<..i})" define h where "h = beukers_integral1" define A :: int where "A = (\i\n. 2 * poly.coeff p i ^ 2)" define B1 where "B1 = (\(i,k)\I. real_of_int (-2 * coeff p i ^ 2) / real_of_int (k ^ 3))" define B2 where "B2 = (\((i,j),k)\J. real_of_int (2 * coeff p i * coeff p j) / real_of_int (k^2*(i-j)))" define d where "d = Lcm {1..n} ^ 3" have [simp]: "h i j = h j i" for i j by (simp add: h_def beukers_integral1_swap) have "beukers_integral2 = (\(x,y)\D. (\(i,j)\{..n}\{..n}. coeff p i * coeff p j * -ln (x*y) / (1-x*y) * x ^ i * y ^ j) \lborel)" unfolding beukers_integral2_def by (subst sum.cartesian_product [symmetric]) (simp add: poly_altdef P_def Shleg_def mult_ac case_prod_unfold p_def sum_distrib_left sum_distrib_right sum_negf sum_divide_distrib) also have "\ = (\(i,j)\{..n}\{..n}. coeff p i * coeff p j * h i j)" unfolding case_prod_unfold proof (subst set_integral_sum) fix ij :: "nat \ nat" have "set_integrable lborel D (\(x,y). real_of_int (coeff p (fst ij) * coeff p (snd ij)) * (-ln (x*y) / (1-x*y) * x ^ fst ij * y ^ snd ij))" unfolding case_prod_unfold using beukers_integral1_integrable[of "fst ij" "snd ij"] by (intro set_integrable_mult_right) (auto simp: D_def case_prod_unfold) thus "set_integrable lborel D (\pa. real_of_int (coeff p (fst ij) * coeff p (snd ij)) * -ln (fst pa * snd pa) / (1 - fst pa * snd pa) * fst pa ^ fst ij * snd pa ^ snd ij)" by (simp add: mult_ac case_prod_unfold) qed (auto simp: beukers_integral1_def h_def case_prod_unfold mult.assoc D_def simp flip: set_integral_mult_right) also have "\ = (\(i,j)\?I1\?I2. coeff p i * coeff p j * h i j)" by (intro sum.cong) auto also have "\ = (\(i,j)\?I1. coeff p i * coeff p j * h i j) + (\(i,j)\?I2. coeff p i * coeff p j * h i j)" by (intro sum.union_disjoint) auto also have "(\(i,j)\?I1. coeff p i * coeff p j * h i j) = (\i\n. coeff p i ^ 2 * h i i)" by (subst sum.reindex) (auto intro: inj_onI simp: case_prod_unfold power2_eq_square) also have "\ = (\i\n. coeff p i ^ 2 * 2 * (Re (zeta 3) - (\k=1..i. 1 / k ^ 3)))" unfolding h_def D_def by (intro sum.cong refl, subst beukers_integral1_same) auto also have "\ = of_int A * Re (zeta 3) - (\i\n. 2 * coeff p i ^ 2 * (\k=1..i. 1 / k ^ 3))" by (simp add: sum_subtractf sum_distrib_left sum_distrib_right algebra_simps A_def) also have "\ = of_int A * Re (zeta 3) + B1" unfolding I_def B1_def by (subst sum.Sigma [symmetric]) (auto simp: sum_distrib_left sum_negf) also have "(\(i,j)\?I2. coeff p i * coeff p j * h i j) = (\(i,j)\?I3\?I4. coeff p i * coeff p j * h i j)" by (intro sum.cong) auto also have "\ = (\(i,j)\?I3. coeff p i * coeff p j * h i j) + (\(i,j)\?I4. coeff p i * coeff p j * h i j)" by (intro sum.union_disjoint) auto also have "(\(i,j)\?I3. coeff p i * coeff p j * h i j) = (\(i,j)\?I4. coeff p i * coeff p j * h i j)" by (intro sum.reindex_bij_witness[of _ "\(i,j). (j,i)" "\(i,j). (j,i)"]) auto also have "\ + \ = 2 * \" by simp also have "\ = (\(i,j)\?I4. \k\{j<..i}. 2 * coeff p i * coeff p j / k ^ 2 / (i - j))" unfolding sum_distrib_left by (intro sum.cong refl) (auto simp: h_def beukers_integral1_different sum_divide_distrib sum_distrib_left mult_ac) also have "\ = B2" unfolding J_def B2_def by (subst sum.Sigma [symmetric]) (auto simp: case_prod_unfold) also have "\B1'. B1 = real_of_int B1' / real_of_int d" unfolding B1_def case_prod_unfold by (rule sum_rationals_common_divisor') (auto simp: d_def I_def) then obtain B1' where "B1 = real_of_int B1' / real_of_int d" .. also have "\B2'. B2 = real_of_int B2' / real_of_int d" unfolding B2_def case_prod_unfold J_def proof (rule sum_rationals_common_divisor'; clarsimp?) fix i j k :: nat assume ijk: "i \ n" "j < k" "k \ i" have "int (k ^ 2 * (i - j)) dvd int (Lcm {1..n} ^ 2 * Lcm {1..n})" unfolding int_dvd_int_iff using ijk by (intro mult_dvd_mono dvd_power_same dvd_Lcm) auto also have "\ = d" by (simp add: d_def power_numeral_reduce) finally show "int k ^ 2 * int (i - j) dvd int d" by simp qed(auto simp: d_def J_def intro!: Nat.gr0I) then obtain B2' where "B2 = real_of_int B2' / real_of_int d" .. finally have "beukers_integral2 = of_int A * Re (zeta 3) + of_int (B1' + B2') / of_nat (Lcm {1..n} ^ 3)" by (simp add: add_divide_distrib d_def) moreover have "coeff p 0 = P 0" unfolding P_def p_def Shleg_def by (simp add: poly_0_coeff_0) hence "coeff p 0 = 1" by (simp add: P_def) hence "A > 0" unfolding A_def by (intro sum_pos2[of _ 0]) auto ultimately show ?thesis by (intro that[of A "B1' + B2'"]) auto qed lemma beukers_integral2_integrable: "set_integrable lborel D (\(x,y). -ln (x*y) / (1 - x*y) * P x * P y)" proof - have "bounded (P ` {0..1})" unfolding P_def Shleg_def by (intro compact_imp_bounded compact_continuous_image continuous_intros) auto then obtain C where C: "\x. x \ {0..1} \ norm (P x) \ C" unfolding bounded_iff by fast have [measurable]: "P \ borel_measurable borel" by (simp add: P_def) from C[of 0] have "C \ 0" by simp show ?thesis proof (rule set_integrable_bound[OF _ _ AE_I2]; clarify?) show "set_integrable lborel D (\(x,y). C ^ 2 * (-ln (x*y) / (1 - x*y)))" using beukers_integral1_integrable[of 0 0] unfolding case_prod_unfold by (intro set_integrable_mult_right) (auto simp: D_def) next fix x y :: real assume xy: "(x, y) \ D" from xy have "x * y < 1" using mult_strict_mono[of x 1 y 1] by (simp add: D_def) have "norm (-ln (x*y) / (1 - x*y) * P x * P y) = (-ln (x*y)) / (1 - x*y) * norm (P x) * norm (P y)" using xy \x * y < 1\ by (simp add: abs_mult abs_divide D_def) also have "\ \ (-ln (x*y)) / (1-x*y) * C * C" using xy C[of x] C[of y] \x * y < 1\ \C \ 0\ by (intro mult_mono divide_left_mono) (auto simp: D_def divide_nonpos_nonneg mult_nonpos_nonneg) also have "\ = norm ((-ln (x*y)) / (1-x*y) * C * C)" using xy \x * y < 1\ \C \ 0\ by (simp add: abs_divide abs_mult D_def) finally show "norm (-ln (x*y) / (1 - x*y) * P x * P y) \ norm (case (x, y) of (x, y) \ C\<^sup>2 * (- ln (x * y) / (1 - x * y)))" by (auto simp: algebra_simps power2_eq_square abs_mult abs_divide) qed (auto simp: D_def set_borel_measurable_def case_prod_unfold simp flip: lborel_prod) qed subsection \The triple integral\ text \ Lastly, we turn to the triple integral \[I_3 := \int_0^1 \int_0^1 \int_0^1 \frac{(x(1-x)y(1-y)w(1-w))^n}{(1-(1-xy)w)^{n+1}}\ \text{d}x\,\text{d}y\,\text{d}w\ .\] \ definition beukers_nn_integral3 :: ennreal where "beukers_nn_integral3 = (\\<^sup>+(w,x,y)\D'. ((x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1)) \lborel)" definition beukers_integral3 :: real where "beukers_integral3 = (\(w,x,y)\D'. ((x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1)) \lborel)" text \ We first prove the following bound (which is a consequence of the arithmetic--geometric mean inequality) that will help us to bound the triple integral. \ lemma beukers_integral3_integrand_bound: fixes x y z :: real assumes xyz: "x \ {0<..<1}" "y \ {0<..<1}" "z \ {0<..<1}" shows "(x*(1-x)*y*(1-y)*z*(1-z)) / (1-(1-x*y)*z) \ 1 / 27" (is "?lhs \ _") proof - have ineq1: "x * (1 - x) \ 1 / 4" if x: "x \ {0..1}" for x :: real proof - have "x * (1 - x) - 1 / 4 = -((x - 1 / 2) ^ 2)" by (simp add: algebra_simps power2_eq_square) also have "\ \ 0" by simp finally show ?thesis by simp qed have ineq2: "x * (1 - x) ^ 2 \ 4 / 27" if x: "x \ {0..1}" for x :: real proof - have "x * (1 - x) ^ 2 - 4 / 27 = (x - 4 / 3) * (x - 1 / 3) ^ 2" by (simp add: algebra_simps power2_eq_square) also have "\ \ 0" by (rule mult_nonpos_nonneg) (use x in auto) finally show ?thesis by simp qed have "1 - (1-x*y)*z = (1 - z) + x * y * z" by (simp add: algebra_simps) also have "\ \ 2 * sqrt (1 - z) * sqrt x * sqrt y * sqrt z" using arith_geo_mean_sqrt[of "1 - z" "x * y * z"] xyz by (auto simp: real_sqrt_mult) finally have *: "?lhs \ (x*(1-x)*y*(1-y)*z*(1-z)) / (2 * sqrt (1 - z) * sqrt x * sqrt y * sqrt z)" using xyz beukers_denom_ineq[of x y z] by (intro divide_left_mono mult_nonneg_nonneg mult_pos_pos) auto have "(x*(1-x)*y*(1-y)*z*(1-z)) = (sqrt x * sqrt x * (1-x) * sqrt y * sqrt y * (1-y) * sqrt z * sqrt z * sqrt (1-z) * sqrt (1-z))" using xyz by simp also have "\ / (2 * sqrt (1 - z) * sqrt x * sqrt y * sqrt z) = sqrt (x * (1 - x) ^ 2) * sqrt (y * (1 - y) ^ 2) * sqrt (z * (1 - z)) / 2" using xyz by (simp add: divide_simps real_sqrt_mult del: real_sqrt_mult_self) also have "\ \ sqrt (4 / 27) * sqrt (4 / 27) * sqrt (1 / 4) / 2" using xyz by (intro divide_right_mono mult_mono real_sqrt_le_mono ineq1 ineq2) auto also have "\ = 1 / 27" by (simp add: real_sqrt_divide) finally show ?thesis using * by argo qed text \ Connecting the above bound with our results of $I_1$, it is easy to see that $I_3 \leq 2 \cdot 27^{-n} \cdot \zeta(3)$: \ lemma beukers_nn_integral3_le: "beukers_nn_integral3 \ ennreal (2 * (1 / 27) ^ n * Re (zeta 3))" proof - have D' [measurable]: "D' \ sets (borel \\<^sub>M borel \\<^sub>M borel)" unfolding D'_def by (simp flip: borel_prod) have "beukers_nn_integral3 = (\\<^sup>+(w,x,y)\D'. ((x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1)) \lborel)" by (simp add: beukers_nn_integral3_def) also have "\ \ (\\<^sup>+(w,x,y)\D'. ((1 / 27) ^ n / (1-(1-x*y)*w)) \lborel)" proof (intro set_nn_integral_mono ennreal_leI, clarify, goal_cases) case (1 w x y) have "(x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1) = ((x*(1-x)*y*(1-y)*w*(1-w)) / (1-(1-x*y)*w)) ^ n / (1-(1-x*y)*w)" by (simp add: divide_simps) also have "\ \ (1 / 27) ^ n / (1 - (1 - x * y) * w)" using beukers_denom_ineq[of x y w] 1 by (intro divide_right_mono power_mono beukers_integral3_integrand_bound) (auto simp: D'_def) finally show ?case . qed also have "\ = ennreal ((1 / 27) ^ n) * (\\<^sup>+(w,x,y)\D'. (1 / (1-(1-x*y)*w)) \lborel)" unfolding lborel_prod [symmetric] case_prod_unfold by (subst nn_integral_cmult [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def simp flip: ennreal_mult') also have "(\\<^sup>+(w,x,y)\D'. (1 / (1-(1-x*y)*w)) \lborel) = (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. ennreal (- (ln (x * y) / (1 - x * y))) \lborel)" using beukers_nn_integral1_altdef[of 0 0] by (simp add: beukers_nn_integral1_def D'_def case_prod_unfold) also have "\ = ennreal (2 * Re (zeta 3))" using beukers_nn_integral1_same[of 0 0] by (simp add: D_def beukers_nn_integral1_def) also have "ennreal ((1 / 27) ^ n) * \ = ennreal (2 * (1 / 27) ^ n * Re (zeta 3))" by (subst ennreal_mult' [symmetric]) (simp_all add: mult_ac) finally show ?thesis . qed lemma beukers_nn_integral3_finite: "beukers_nn_integral3 < \" by (rule le_less_trans, rule beukers_nn_integral3_le) simp_all lemma beukers_integral3_integrable: "set_integrable lborel D' (\(w,x,y). (x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1))" unfolding case_prod_unfold using less_imp_le[OF beukers_denom_ineq] beukers_nn_integral3_finite by (intro set_integrableI_nonneg AE_I2 impI) (auto simp: D'_def set_borel_measurable_def beukers_nn_integral3_def case_prod_unfold simp flip: lborel_prod intro!: divide_nonneg_nonneg mult_nonneg_nonneg) lemma beukers_integral3_conv_nn_integral: "beukers_integral3 = enn2real beukers_nn_integral3" unfolding beukers_integral3_def using beukers_nn_integral3_finite less_imp_le[OF beukers_denom_ineq] by (intro set_integral_eq_nn_integral AE_I2 impI) (auto simp: D'_def set_borel_measurable_def beukers_nn_integral3_def case_prod_unfold simp flip: lborel_prod) lemma beukers_integral3_le: "beukers_integral3 \ 2 * (1 / 27) ^ n * Re (zeta 3)" proof - have "beukers_integral3 = enn2real beukers_nn_integral3" by (rule beukers_integral3_conv_nn_integral) also have "\ \ enn2real (ennreal (2 * (1 / 27) ^ n * Re (zeta 3)))" by (intro enn2real_mono beukers_nn_integral3_le) auto also have "\ = 2 * (1 / 27) ^ n * Re (zeta 3)" using Re_zeta_ge_1[of 3] by (intro enn2real_ennreal mult_nonneg_nonneg) auto finally show ?thesis . qed text \ It is also easy to see that $I_3 > 0$. \ lemma beukers_nn_integral3_pos: "beukers_nn_integral3 > 0" proof - have D' [measurable]: "D' \ sets (borel \\<^sub>M borel \\<^sub>M borel)" unfolding D'_def by (simp flip: borel_prod) have *: "\(AE (w,x,y) in lborel. ennreal ((x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1)) * indicator D' (w,x,y) \ 0)" (is "\(AE z in lborel. ?P z)") proof - { fix w x y :: real assume xyw: "(w,x,y) \ D'" hence "(x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1) > 0" using beukers_denom_ineq[of x y w] by (intro divide_pos_pos mult_pos_pos zero_less_power) (auto simp: D'_def) with xyw have "\?P (w,x,y)" by (auto simp: indicator_def D'_def) } hence *: "\?P z" if "z \ D'" for z using that by blast hence "{z\space lborel. \?P z} = D'" by auto moreover have "emeasure lborel D' = 1" proof - have "D' = box (0,0,0) (1,1,1)" by (auto simp: D'_def box_def Basis_prod_def) also have "emeasure lborel \ = 1" by (subst emeasure_lborel_box) (auto simp: Basis_prod_def) finally show ?thesis by simp qed ultimately show ?thesis by (subst AE_iff_measurable[of D']) (simp_all flip: borel_prod) qed hence "nn_integral lborel (\_::real\real\real. 0) < beukers_nn_integral3" unfolding beukers_nn_integral3_def by (intro nn_integral_less) (simp_all add: case_prod_unfold flip: lborel_prod) thus ?thesis by simp qed lemma beukers_integral3_pos: "beukers_integral3 > 0" proof - have "0 < enn2real beukers_nn_integral3" using beukers_nn_integral3_pos beukers_nn_integral3_finite by (subst enn2real_positive_iff) auto also have "\ = beukers_integral3" by (rule beukers_integral3_conv_nn_integral [symmetric]) finally show ?thesis . qed subsection \Connecting the double and triple integral\ text \ In this section, we will prove the most technically involved part, namely that $I_2 = I_3$. I will not go into detail about how this works -- the reader is advised to simply look at Filaseta's presentation of the proof. The basic idea is to integrate by parts \n\ times with respect to \y\ to eliminate the factor $P(y)$, then change variables $z = \frac{1-w}{1-(1-xy)w}$, and then apply the same integration by parts \n\ times to \x\ to eliminate $P(x)$. \ text \ The first expand \[-\frac{\ln (xy)}{1-xy} = \int_0^1 \frac{1}{1-(1-xy)z}\,\text{d}z\ .\] \ lemma beukers_aux_ln_conv_integral: fixes x y :: real assumes xy: "x \ {0<..<1}" "y \ {0<..<1}" shows "-ln (x*y) / (1-x*y) = (LBINT z=0..1. 1 / (1-(1-x*y)*z))" proof - have "x * y < 1" using mult_strict_mono[of x 1 y 1] xy by simp have less: "(1 - x * y) * u < 1" if u: "u \ {0..1}" for u proof - from u \x * y < 1\ have "(1 - x * y) * u \ (1 - x * y) * 1" by (intro mult_left_mono) auto also have "\ < 1 * 1" using xy by (intro mult_strict_right_mono) auto finally show "(1 - x * y) * u < 1" by simp qed have neq: "(1 - x * y) * u \ 1" if "u \ {0..1}" for u using less[of u] that by simp let ?F = "\z. ln (1-(1-x*y)*z)/(x*y-1)" have "(LBINT z=ereal 0..ereal 1. 1 / (1-(1-x*y)*z)) = ?F 1 - ?F 0" proof (rule interval_integral_FTC_finite, goal_cases cont deriv) case cont show ?case using neq by (intro continuous_intros) auto next case (deriv z) show ?case unfolding has_real_derivative_iff_has_vector_derivative [symmetric] by (insert less[of z] xy \x * y < 1\ deriv) (rule derivative_eq_intros refl | simp)+ qed also have "\ = -ln (x*y) / (1-x*y)" using \x * y < 1\ by (simp add: field_simps) finally show ?thesis by (simp add: zero_ereal_def one_ereal_def) qed text \ The first part we shall show is the integration by parts. \ lemma beukers_aux_by_parts_aux: assumes xz: "x \ {0<..<1}" "z \ {0<..<1}" and "k \ n" shows "(LBINT y=0..1. Q n y * (1/(1-(1-x*y)*z))) = (LBINT y=0..1. Q (n-k) y * (fact k * (x*z)^k / (1-(1-x*y)*z) ^ (k+1)))" using assms(3) proof (induction k) case (Suc k) note [derivative_intros] = DERIV_chain2[OF has_field_derivative_Gen_Shleg] define G where "G = (\y. -fact k * (x*z)^k / (1-(1-x*y)*z) ^ (k+1))" define g where "g = (\y. fact (Suc k) * (x*z)^Suc k / (1-(1-x*y)*z) ^ (k+2))" have less: "(1 - x * y) * z < 1" and neq: "(1 - x * y) * z \ 1" if y: "y \ {0..1}" for y proof - from y xz have "x * y \ x * 1" by (intro mult_left_mono) auto also have "\ < 1" using xz by simp finally have "(1 - x * y) * z \ 1 * z" using xz y by (intro mult_right_mono) auto also have "\ < 1" using xz by simp finally show "(1 - x * y) * z < 1" by simp thus "(1 - x * y) * z \ 1" by simp qed have cont: "continuous_on {0..1} g" using neq by (auto simp: g_def intro!: continuous_intros) have deriv: "(G has_real_derivative g y) (at y within {0..1})" if y: "y \ {0..1}" for y unfolding G_def by (insert neq xz y, (rule derivative_eq_intros refl power_not_zero)+) (auto simp: divide_simps g_def) have deriv2: "(Q (n - Suc k) has_real_derivative Q (n - k) y) (at y within {0..1})" for y using Suc.prems by (auto intro!: derivative_eq_intros simp: Suc_diff_Suc Q_def) have "(LBINT y=0..1. Q (n-Suc k) y * (fact (Suc k) * (x*z)^Suc k / (1-(1-x*y)*z) ^ (k+2))) = (LBINT y=0..1. Q (n-Suc k) y * g y)" by (simp add: g_def) also have "(LBINT y=0..1. Q (n-Suc k) y * g y) = -(LBINT y=0..1. Q (n-k) y * G y)" using Suc.prems deriv deriv2 cont by (subst interval_lebesgue_integral_by_parts_01[where f = "Q (n-k)" and G = G]) (auto intro!: continuous_intros simp: Q_def) also have "\ = (LBINT y=0..1. Q (n-k) y * (fact k * (x*z)^k / (1-(1-x*y)*z) ^ (k+1)))" by (simp add: G_def flip: interval_lebesgue_integral_uminus) finally show ?case using Suc by simp qed auto lemma beukers_aux_by_parts: assumes xz: "x \ {0<..<1}" "z \ {0<..<1}" shows "(LBINT y=0..1. P y / (1-(1-x*y)*z)) = (LBINT y=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z) ^ (n+1))" proof - have "(LBINT y=0..1. P y * (1/(1-(1-x*y)*z))) = 1 / fact n * (LBINT y=0..1. Q n y * (1/(1-(1-x*y)*z)))" unfolding interval_lebesgue_integral_mult_right [symmetric] by (simp add: P_def Q_def Shleg_altdef) also have "\ = (LBINT y=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z) ^ (n+1))" by (subst beukers_aux_by_parts_aux[OF assms, of n], simp, subst interval_lebesgue_integral_mult_right [symmetric]) (simp add: Q_def mult_ac Gen_Shleg_0_left power_mult_distrib) finally show ?thesis by simp qed text \ We then get the following by applying the integration by parts to \y\: \ lemma beukers_aux_integral_transform1: fixes z :: real assumes z: "z \ {0<..<1}" shows "(LBINT (x,y):D. P x * P y / (1-(1-x*y)*z)) = (LBINT (x,y):D. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" proof - have cbox: "cbox (0, 0) (1, 1) = ({0..1} \ {0..1} :: (real \ real) set)" by (auto simp: cbox_def Basis_prod_def inner_prod_def) have box: "box (0, 0) (1, 1) = ({0<..<1} \ {0<..<1} :: (real \ real) set)" by (auto simp: box_def Basis_prod_def inner_prod_def) have "set_integrable lborel (cbox (0,0) (1,1)) (\(x, y). P x * P y / (1 - (1 - x * y) * z))" unfolding lborel_prod case_prod_unfold P_def proof (intro continuous_on_imp_set_integrable_cbox continuous_intros ballI) fix p :: "real \ real" assume p: "p \ cbox (0, 0) (1, 1)" have "(1 - fst p * snd p) * z \ 1 * z" using mult_mono[of "fst p" 1 "snd p" 1] p z cbox by (intro mult_right_mono) auto also have "\ < 1" using z by simp finally show "1 - (1 - fst p * snd p) * z \ 0" by simp qed hence integrable: "set_integrable lborel (box (0,0) (1,1)) (\(x, y). P x * P y / (1 - (1 - x * y) * z))" by (rule set_integrable_subset) (auto simp: box simp flip: borel_prod) have "set_integrable lborel (cbox (0,0) (1,1)) (\(x, y). P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" unfolding lborel_prod case_prod_unfold P_def proof (intro continuous_on_imp_set_integrable_cbox continuous_intros ballI) fix p :: "real \ real" assume p: "p \ cbox (0, 0) (1, 1)" have "(1 - fst p * snd p) * z \ 1 * z" using mult_mono[of "fst p" 1 "snd p" 1] p z cbox by (intro mult_right_mono) auto also have "\ < 1" using z by simp finally show "(1 - (1 - fst p * snd p) * z) ^ (n + 1) \ 0" by simp qed hence integrable': "set_integrable lborel D (\(x, y). P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" by (rule set_integrable_subset) (auto simp: box D_def simp flip: borel_prod) have "(LBINT (x,y):D. P x * P y / (1-(1-x*y)*z)) = (LBINT x=0..1. (LBINT y=0..1. P x * P y / (1-(1-x*y)*z)))" unfolding D_def lborel_prod [symmetric] using box integrable by (subst lborel_pair.set_integral_fst') (simp_all add: interval_integral_Ioo lborel_prod) also have "\ = (LBINT x=0..1. P x * (LBINT y=0..1. P y / (1-(1-x*y)*z)))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (simp add: mult_ac) also have "\ = (LBINT x=0..1. P x * (LBINT y=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)))" using z by (intro interval_lebesgue_integral_lborel_01_cong, subst beukers_aux_by_parts) auto also have "\ = (LBINT x=0..1. (LBINT y=0..1. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (simp add: mult_ac) also have "\ = (LBINT (x,y):D. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" unfolding D_def lborel_prod [symmetric] using box integrable' by (subst lborel_pair.set_integral_fst') (simp_all add: D_def interval_integral_Ioo lborel_prod) finally show "(LBINT (x,y):D. P x * P y / (1-(1-x*y)*z)) = \" . qed text \ We then change variables for \z\: \ lemma beukers_aux_integral_transform2: assumes xy: "x \ {0<..<1}" "y \ {0<..<1}" shows "(LBINT z=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)) = (LBINT w=0..1. (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" proof - define g where "g = (\z. (1 - z) / (1-(1-x*y)*z))" define g' where "g' = (\z. -x*y / (1-(1-x*y)*z)^2)" have "x * y < 1" using mult_strict_mono[of x 1 y 1] xy by simp have less: "(1 - (x * y)) * w < 1" and neq: "(1 - (x * y)) * w \ 1" if w: "w \ {0..1}" for w proof - have "(1 - (x * y)) * w \ (1 - (x * y)) * 1" using w \x * y < 1\ by (intro mult_left_mono) auto also have "\ < 1" using xy by simp finally show "(1 - (x * y)) * w < 1" . thus "(1 - (x * y)) * w \ 1" by simp qed have deriv: "(g has_real_derivative g' w) (at w within {0..1})" if "w \ {0..1}" for w unfolding g_def g'_def apply (insert that xy neq) apply (rule derivative_eq_intros refl)+ apply (simp_all add: divide_simps power2_eq_square) apply (auto simp: algebra_simps) done have "continuous_on {0..1} (\xa. (1 - xa) ^ n / (1 - (1 - x * y) * xa))" using neq by (auto intro!: continuous_intros) moreover have "g ` {0..1} \ {0..1}" proof clarify fix w :: real assume w: "w \ {0..1}" have "(1 - x * y) * w \ 1 * w" using \x * y < 1\ xy w by (intro mult_right_mono) auto thus "g w \ {0..1}" unfolding g_def using less[of w] w by (auto simp: divide_simps) qed ultimately have cont: "continuous_on (g ` {0..1}) (\xa. (1 - xa) ^ n / (1 - (1 - x * y) * xa))" by (rule continuous_on_subset) have cont': "continuous_on {0..1} g'" using neq by (auto simp: g'_def intro!: continuous_intros) have "(LBINT w=0..1. (1-w)^n * (1-y)^n / (1-(1-x*y)*w)) = (1-y)^n * (LBINT w=0..1. (1 - w)^n / (1-(1-x*y)*w))" unfolding interval_lebesgue_integral_mult_right [symmetric] by (simp add: algebra_simps power_mult_distrib) also have "(LBINT w=0..1. (1 - w)^n / (1-(1-x*y)*w)) = -(LBINT w=g 0..g 1. (1 - w)^n / (1-(1-x*y)*w))" by (subst interval_integral_endpoints_reverse)(simp add: g_def zero_ereal_def one_ereal_def) also have "(LBINT w=g 0..g 1. (1 - w)^n / (1-(1-x*y)*w)) = (LBINT w=0..1. g' w * ((1 - g w) ^ n / (1 - (1-x*y) * g w)))" using deriv cont cont' by (subst interval_integral_substitution_finite [symmetric, where g = g and g' = g']) (simp_all add: zero_ereal_def one_ereal_def) also have "-\ = (LBINT z=0..1. ((x*y)^n * z^n / (1-(1-x*y)*z)^(n+1)))" unfolding interval_lebesgue_integral_uminus [symmetric] using xy apply (intro interval_lebesgue_integral_lborel_01_cong) apply (simp add: divide_simps g_def g'_def) apply (auto simp: algebra_simps power_mult_distrib power2_eq_square) done also have "(1-y)^n * \ = (LBINT z=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" unfolding interval_lebesgue_integral_mult_right [symmetric] by (simp add: algebra_simps power_mult_distrib) finally show "\ = (LBINT w=0..1. (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" .. qed text \ Lastly, we apply the same integration by parts to \x\: \ lemma beukers_aux_integral_transform3: assumes w: "w \ {0<..<1}" shows "(LBINT (x,y):D. P x * (1-y)^n / (1-(1-x*y)*w)) = (LBINT (x,y):D. (1-y)^n * (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1))" proof - have cbox: "cbox (0, 0) (1, 1) = ({0..1} \ {0..1} :: (real \ real) set)" by (auto simp: cbox_def Basis_prod_def inner_prod_def) have box: "box (0, 0) (1, 1) = ({0<..<1} \ {0<..<1} :: (real \ real) set)" by (auto simp: box_def Basis_prod_def inner_prod_def) have "set_integrable lborel (cbox (0,0) (1,1)) (\(x,y). P x * (1-y)^n/(1-(1-x*y)*w))" unfolding lborel_prod case_prod_unfold P_def proof (intro continuous_on_imp_set_integrable_cbox continuous_intros ballI) fix p :: "real \ real" assume p: "p \ cbox (0,0) (1,1)" have "(1 - fst p * snd p) * w \ 1 * w" using p cbox w by (intro mult_right_mono) auto also have "\ < 1" using w by simp finally have "(1 - fst p * snd p) * w < 1" by simp thus "1 - (1 - fst p * snd p) * w \ 0" by simp qed hence integrable: "set_integrable lborel D (\(x,y). P x * (1-y)^n/(1-(1-x*y)*w))" by (rule set_integrable_subset) (auto simp: D_def simp flip: borel_prod) have "set_integrable lborel (cbox (0,0) (1,1)) (\(x,y). (1-y)^n * (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1))" unfolding lborel_prod case_prod_unfold P_def proof (intro continuous_on_imp_set_integrable_cbox continuous_intros ballI) fix p :: "real \ real" assume p: "p \ cbox (0,0) (1,1)" have "(1 - fst p * snd p) * w \ 1 * w" using p cbox w by (intro mult_right_mono) auto also have "\ < 1" using w by simp finally have "(1 - fst p * snd p) * w < 1" by simp thus "(1 - (1 - fst p * snd p) * w) ^ (n + 1) \ 0" by simp qed hence integrable': "set_integrable lborel D (\(x,y). (1-y)^n * (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1))" by (rule set_integrable_subset) (auto simp: D_def simp flip: borel_prod) have "(LBINT (x,y):D. P x * (1-y)^n / (1-(1-x*y)*w)) = (LBINT y=0..1. (LBINT x=0..1. P x * (1-y)^n / (1-(1-x*y)*w)))" using integrable unfolding case_prod_unfold D_def lborel_prod [symmetric] by (subst lborel_pair.set_integral_snd) (auto simp: interval_integral_Ioo) also have "\ = (LBINT y=0..1. (1-y)^n * (LBINT x=0..1. P x / (1-(1-y*x)*w)))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (auto simp: mult_ac) also have "\ = (LBINT y=0..1. (1-y)^n * (LBINT x=0..1. (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1)))" using w by (intro interval_lebesgue_integral_lborel_01_cong, subst beukers_aux_by_parts) (auto simp: mult_ac) also have "\ = (LBINT y=0..1. (LBINT x=0..1. (1-y)^n * (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1)))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (auto simp: mult_ac) also have "\ = (LBINT (x,y):D. (1-y)^n * (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1))" using integrable' unfolding case_prod_unfold D_def lborel_prod [symmetric] by (subst lborel_pair.set_integral_snd) (auto simp: interval_integral_Ioo) finally show ?thesis . qed text \ We need to show the existence of some of these triple integrals. \ lemma beukers_aux_integrable1: "set_integrable lborel (({0<..<1} \ {0<..<1}) \ {0<..<1}) (\((x,y),z). P x * P y / (1-(1-x*y)*z))" proof - have D [measurable]: "D \ sets (borel \\<^sub>M borel)" unfolding D_def by (simp flip: borel_prod) have "bounded (P ` {0..1})" unfolding P_def by (intro compact_imp_bounded compact_continuous_image continuous_intros) auto then obtain C where C: "\x. x \ {0..1} \ norm (P x) \ C" unfolding bounded_iff by fast show ?thesis unfolding D'_def case_prod_unfold proof (subst lborel_prod [symmetric], intro lborel_pair.Fubini_set_integrable AE_I2 impI; clarsimp?) fix x y :: real assume xy: "x > 0" "x < 1" "y > 0" "y < 1" have "x * y < 1" using xy mult_strict_mono[of x 1 y 1] by simp show "set_integrable lborel {0<..<1} (\z. P x * P y / (1-(1-x*y)*z))" by (rule set_integrable_subset[of _ "{0..1}"], rule borel_integrable_atLeastAtMost') (use \x*y<1\ beukers_denom_neq[of x y] xy in \auto intro!: continuous_intros simp: P_def\) next have "set_integrable lborel D (\(x,y). (\z\{0<..<1}. norm (P x * P y / (1-(1-x*y)*z)) \lborel))" proof (rule set_integrable_bound[OF _ _ AE_I2]; clarify?) show "set_integrable lborel D (\(x,y). C\<^sup>2 * (-ln (x*y) / (1 - x*y)))" using beukers_integral1_integrable[of 0 0] unfolding case_prod_unfold by (intro set_integrable_mult_right) (auto simp: D_def) next fix x y assume xy: "(x, y) \ D" have "norm (LBINT z:{0<..<1}. norm (P x * P y / (1-(1-x*y)*z))) = norm (LBINT z:{0<..<1}. \P x\ * \P y\ * (1 / (1-(1-x*y)*z)))" proof (intro arg_cong[where f = norm] set_lebesgue_integral_cong allI impI, goal_cases) case (2 z) with beukers_denom_ineq[of x y z] xy show ?case by (auto simp: abs_mult D_def) qed (auto simp: abs_mult D_def) also have "\ = norm (\P x\ * \P y\ * (LBINT z=0..1. (1 / (1-(1-x*y)*z))))" by (subst set_integral_mult_right) (auto simp: interval_integral_Ioo) also have "\ = norm (norm (P x) * norm (P y) * (- ln (x * y) / (1 - x * y)))" using beukers_aux_ln_conv_integral[of x y] xy by (simp add: D_def) also have "\ = norm (P x) * norm (P y) * (- ln (x * y) / (1 - x * y))" using xy mult_strict_mono[of x 1 y 1] by (auto simp: D_def divide_nonpos_nonneg abs_mult) also have "norm (P x) * norm (P y) * (- ln (x * y) / (1 - x * y)) \ norm (C * C * (- ln (x * y) / (1 - x * y)))" using xy C[of x] C[of y] mult_strict_mono[of x 1 y 1] unfolding norm_mult norm_divide by (intro mult_mono C) (auto simp: D_def divide_nonpos_nonneg) finally show "norm (LBINT z:{0<..<1}. norm (P x * P y / (1-(1-x*y)*z))) \ norm (case (x, y) of (x, y) \ C\<^sup>2 * (- ln (x * y) / (1 - x * y)))" by (simp add: power2_eq_square mult_ac) next show "set_borel_measurable lborel D (\(x, y). LBINT z:{0<..<1}. norm (P x * P y / (1 - (1 - x * y) * z)))" unfolding lborel_prod [symmetric] set_borel_measurable_def case_prod_unfold set_lebesgue_integral_def P_def by measurable qed thus "set_integrable lborel ({0<..<1} \ {0<..<1}) (\x. LBINT y:{0<..<1}. \P (fst x) * P (snd x)\ / \1 - (1 - fst x * snd x) * y\)" by (simp add: case_prod_unfold D_def) qed (auto simp: case_prod_unfold lborel_prod [symmetric] set_borel_measurable_def P_def) qed lemma beukers_aux_integrable2: "set_integrable lborel D' (\(z,x,y). P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z) ^ (n+1))" proof - have [measurable]: "P \ borel_measurable borel" unfolding P_def by (intro borel_measurable_continuous_onI continuous_intros) have "bounded (P ` {0..1})" unfolding P_def by (intro compact_imp_bounded compact_continuous_image continuous_intros) auto then obtain C where C: "\x. x \ {0..1} \ norm (P x) \ C" unfolding bounded_iff by fast show ?thesis unfolding D'_def proof (rule set_integrable_bound[OF _ _ AE_I2]; clarify?) show "set_integrable lborel ({0<..<1} \ {0<..<1} \ {0<..<1}) (\(z,x,y). C * (1 / (1-(1-x*y)*z)))" unfolding case_prod_unfold using beukers_integral1_integrable'[of 0 0] by (intro set_integrable_mult_right) (auto simp: lborel_prod case_prod_unfold) next fix x y z :: real assume xyz: "x \ {0<..<1}" "y \ {0<..<1}" "z \ {0<..<1}" have "norm (P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)) = norm (P x) * (1-y)^n * ((x*y*z) / (1-(1-x*y)*z))^n / (1-(1-x*y)*z)" using xyz beukers_denom_ineq[of x y z] by (simp add: abs_mult power_divide mult_ac) also have "(x*y*z) / (1-(1-x*y)*z) = 1/((1-z)/(z*x*y)+1)" using xyz by (simp add: field_simps) also have "norm (P x) * (1-y)^n * \^n / (1-(1-x*y)*z) \ C * 1^n * 1^n / (1-(1-x*y)*z)" using xyz C[of x] beukers_denom_ineq[of x y z] by (intro mult_mono divide_right_mono power_mono zero_le_power mult_nonneg_nonneg divide_nonneg_nonneg) (auto simp: field_simps) also have "\ \ \C * 1^n * 1^n / (1-(1-x*y)*z)\" by linarith finally show "norm (P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)) \ norm (case (z,x,y) of (z,x,y) \ C * (1 / (1-(1-x*y)*z)))" by (simp add: case_prod_unfold) qed (simp_all add: case_prod_unfold set_borel_measurable_def flip: borel_prod) qed lemma beukers_aux_integrable3: "set_integrable lborel D' (\(w,x,y). P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" proof - have [measurable]: "P \ borel_measurable borel" unfolding P_def by (intro borel_measurable_continuous_onI continuous_intros) have "bounded (P ` {0..1})" unfolding P_def by (intro compact_imp_bounded compact_continuous_image continuous_intros) auto then obtain C where C: "\x. x \ {0..1} \ norm (P x) \ C" unfolding bounded_iff by fast show ?thesis unfolding D'_def proof (rule set_integrable_bound[OF _ _ AE_I2]; clarify?) show "set_integrable lborel ({0<..<1} \ {0<..<1} \ {0<..<1}) (\(z,x,y). C * (1 / (1-(1-x*y)*z)))" unfolding case_prod_unfold using beukers_integral1_integrable'[of 0 0] by (intro set_integrable_mult_right) (auto simp: lborel_prod case_prod_unfold) next fix x y w :: real assume xyw: "x \ {0<..<1}" "y \ {0<..<1}" "w \ {0<..<1}" have "norm (P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w)) = norm (P x) * (1-w)^n * (1-y)^n / (1-(1-x*y)*w)" using xyw beukers_denom_ineq[of x y w] by (simp add: abs_mult power_divide mult_ac) also have "\ \ C * 1^n * 1^n / (1-(1-x*y)*w)" using xyw C[of x] beukers_denom_ineq[of x y w] by (intro mult_mono divide_right_mono power_mono zero_le_power mult_nonneg_nonneg divide_nonneg_nonneg) (auto simp: field_simps) also have "\ \ \C * 1^n * 1^n / (1-(1-x*y)*w)\" by linarith finally show "norm (P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w)) \ norm (case (w,x,y) of (z,x,y) \ C * (1 / (1-(1-x*y)*z)))" by (simp add: case_prod_unfold) qed (simp_all add: case_prod_unfold set_borel_measurable_def flip: borel_prod) qed text \ Now we only need to put all of these results together: \ lemma beukers_integral2_conv_3: "beukers_integral2 = beukers_integral3" proof - have cont_P: "continuous_on {0..1} P" by (auto simp: P_def intro: continuous_intros) have D [measurable]: "D \ sets borel" "D \ sets (borel \\<^sub>M borel)" unfolding D_def by (simp_all flip: borel_prod) have [measurable]: "P \ borel_measurable borel" unfolding P_def by (intro borel_measurable_continuous_onI continuous_intros) have "beukers_integral2 = (LBINT (x,y):D. P x * P y * (LBINT z=0..1. 1 / (1-(1-x*y)*z)))" unfolding beukers_integral2_def case_prod_unfold by (intro set_lebesgue_integral_cong allI impI, measurable) (subst beukers_aux_ln_conv_integral, auto simp: D_def) also have "\ = (LBINT (x,y):D. (LBINT z=0..1. P x * P y / (1-(1-x*y)*z)))" by (subst interval_lebesgue_integral_mult_right [symmetric]) auto also have "\ = (LBINT (x,y):D. (LBINT z:{0<..<1}. P x * P y / (1-(1-x*y)*z)))" by (simp add: interval_integral_Ioo) also have "\ = (LBINT z:{0<..<1}. (LBINT (x,y):D. P x * P y / (1-(1-x*y)*z)))" proof (subst lborel_pair.Fubini_set_integral [symmetric]) have "set_integrable lborel (({0<..<1} \ {0<..<1}) \ {0<..<1}) (\((x, y), z). P x * P y / (1 - (1 - x * y) * z))" using beukers_aux_integrable1 by simp also have "?this \ set_integrable (lborel \\<^sub>M lborel) ({0<..<1} \ D) (\(z,x,y). P x * P y / (1 - (1 - x * y) * z))" unfolding set_integrable_def by (subst lborel_pair.integrable_product_swap_iff [symmetric], intro Bochner_Integration.integrable_cong) (auto simp: indicator_def case_prod_unfold lborel_prod D_def) finally show \ . qed (auto simp: case_prod_unfold) also have "\ = (LBINT z:{0<..<1}. (LBINT (x,y):D. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)))" by (rule set_lebesgue_integral_cong) (use beukers_aux_integral_transform1 in simp_all) also have "\ = (LBINT (x,y):D. (LBINT z:{0<..<1}. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)))" using beukers_aux_integrable2 by (subst lborel_pair.Fubini_set_integral [symmetric]) (auto simp: case_prod_unfold lborel_prod D_def D'_def) also have "\ = (LBINT (x,y):D. (LBINT w:{0<..<1}. P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w)))" proof (intro set_lebesgue_integral_cong allI impI; clarify?) fix x y :: real assume xy: "(x, y) \ D" have "(LBINT z:{0<..<1}. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)) = P x * (LBINT z=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (simp add: mult_ac interval_integral_Ioo) also have "\ = P x * (LBINT w=0..1. (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" using xy by (subst beukers_aux_integral_transform2) (auto simp: D_def) also have "\ = (LBINT w:{0<..<1}. P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (simp add: mult_ac interval_integral_Ioo) finally show "(LBINT z:{0<..<1}. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)) = (LBINT w:{0<..<1}. P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" . qed (auto simp: D_def simp flip: borel_prod) also have "\ = (LBINT w:{0<..<1}. (LBINT (x,y):D. P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w)))" using beukers_aux_integrable3 by (subst lborel_pair.Fubini_set_integral [symmetric]) (auto simp: case_prod_unfold lborel_prod D_def D'_def) also have "\ = (LBINT w=0..1. (1-w)^n * (LBINT (x,y):D. P x * (1-y)^n / (1-(1-x*y)*w)))" unfolding case_prod_unfold by (subst set_integral_mult_right [symmetric]) (simp add: mult_ac interval_integral_Ioo) also have "\ = (LBINT w=0..1. (1-w)^n * (LBINT (x,y):D. (x*y*w*(1-x)*(1-y))^n / (1-(1-x*y)*w)^(n+1)))" by (intro interval_lebesgue_integral_lborel_01_cong, subst beukers_aux_integral_transform3) (auto simp: mult_ac power_mult_distrib) also have "\ = (LBINT w=0..1. (LBINT (x,y):D. (x*y*w*(1-x)*(1-y)*(1-w))^n / (1-(1-x*y)*w)^(n+1)))" by (subst set_integral_mult_right [symmetric]) (auto simp: case_prod_unfold mult_ac power_mult_distrib) also have "\ = beukers_integral3" using beukers_integral3_integrable unfolding D'_def D_def beukers_integral3_def by (subst (2) lborel_prod [symmetric], subst lborel_pair.set_integral_fst') (auto simp: case_prod_unfold interval_integral_Ioo lborel_prod algebra_simps) finally show ?thesis . qed subsection \The main result\ text \ Combining all of the results so far, we can derive the key inequalities \[0 < A\zeta(3) + B < 2 \zeta(3) \cdot 27^{-n} \cdot \text{lcm}\{1\ldots n\}^3\] for integers $A$, $B$ with $A > 0$. \ lemma zeta_3_linear_combination_bounds: obtains A B :: int where "A > 0" "A * Re (zeta 3) + B \ {0 <.. 2 * Re (zeta 3) * Lcm {1..n} ^ 3 / 27 ^ n}" proof - define I where "I = beukers_integral2" define d where "d = Lcm {1..n} ^ 3" have "d > 0" by (auto simp: d_def intro!: Nat.gr0I) from beukers_integral2_conv_int_combination obtain A' B :: int where *: "A' > 0" "I = A' * Re (zeta 3) + B / d" unfolding I_def d_def . define A where "A = A' * d" from * have A: "A > 0" "I = (A * Re (zeta 3) + B) / d" using \d > 0\ by (simp_all add: A_def field_simps) have "0 < I" using beukers_integral3_pos by (simp add: I_def beukers_integral2_conv_3) with \d > 0\ have "A * Re (zeta 3) + B > 0" by (simp add: field_simps A) moreover have "I \ 2 * (1 / 27) ^ n * Re (zeta 3)" using beukers_integral2_conv_3 beukers_integral3_le by (simp add: I_def) hence "A * Re (zeta 3) + B \ 2 * Re (zeta 3) * d / 27 ^ n" using \d > 0\ by (simp add: A field_simps) ultimately show ?thesis using A by (intro that[of A B]) (auto simp: d_def) qed text \ If $\zeta(3) = \frac{a}{b}$ for some integers $a$ and $b$, we can thus derive the inequality $2b\zeta(3) \cdot 27^{-n} \cdot \text{lcm}\{1\ldots n\}^3\geq 1$ for any natural number $n$. \ lemma beukers_key_inequality: fixes a :: int and b :: nat assumes "b > 0" and ab: "Re (zeta 3) = a / b" shows "2 * b * Re (zeta 3) * Lcm {1..n} ^ 3 / 27 ^ n \ 1" proof - from zeta_3_linear_combination_bounds obtain A B :: int where AB: "A > 0" "A * Re (zeta 3) + B \ {0 <.. 2 * Re (zeta 3) * Lcm {1..n} ^ 3 / 27 ^ n}" . from AB have "0 < (A * Re (zeta 3) + B) * b" using \b > 0\ by (intro mult_pos_pos) auto also have "\ = A * (Re (zeta 3) * b) + B * b" using \b > 0\ by (simp add: algebra_simps) also have "Re (zeta 3) * b = a" using \b > 0\ by (simp add: ab) also have "of_int A * of_int a + of_int (B * b) = of_int (A * a + B * b)" by simp finally have "1 \ A * a + B * b" by linarith hence "1 \ real_of_int (A * a + B * b)" by linarith also have "\ = (A * (a / b) + B) * b" using \b > 0\ by (simp add: ring_distribs) also have "a / b = Re (zeta 3)" by (simp add: ab) also have "A * Re (zeta 3) + B \ 2 * Re (zeta 3) * Lcm {1..n} ^ 3 / 27 ^ n" using AB by simp finally show "2 * b * Re (zeta 3) * Lcm {1..n} ^ 3 / 27 ^ n \ 1" using \b > 0\ by (simp add: mult_ac) qed end (* TODO: Move *) lemma smallo_power: "n > 0 \ f \ o[F](g) \ (\x. f x ^ n) \ o[F](\x. g x ^ n)" by (induction n rule: nat_induct_non_zero) (auto intro: landau_o.small.mult) text \ This is now a contradiction, since $\text{lcm}\{1\ldots n\} \in o(3^n)$ by the Prime Number Theorem -- hence the main result. \ theorem zeta_3_irrational: "zeta 3 \ \" proof assume "zeta 3 \ \" obtain a :: int and b :: nat where "b > 0" and ab': "zeta 3 = a / b" proof - from \zeta 3 \ \\ obtain r where r: "zeta 3 = of_rat r" by (elim Rats_cases) also have "r = rat_of_int (fst (quotient_of r)) / rat_of_int (snd (quotient_of r))" by (intro quotient_of_div) auto also have "of_rat \ = (of_int (fst (quotient_of r)) / of_int (snd (quotient_of r)) :: complex)" by (simp add: of_rat_divide) also have "of_int (snd (quotient_of r)) = of_nat (nat (snd (quotient_of r)))" using quotient_of_denom_pos'[of r] by auto finally have "zeta 3 = of_int (fst (quotient_of r)) / of_nat (nat (snd (quotient_of r)))" . thus ?thesis using quotient_of_denom_pos'[of r] by (intro that[of "nat (snd (quotient_of r))" "fst (quotient_of r)"]) auto qed hence ab: "Re (zeta 3) = a / b" by simp interpret prime_number_theorem by standard (rule prime_number_theorem) have Lcm_upto_smallo: "(\n. real (Lcm {1..n})) \ o(\n. c ^ n)" if c: "c > exp 1" for c proof - have "0 < exp (1::real)" by simp also note c finally have "c > 0" . have "(\n. real (Lcm {1..n})) = (\n. real (Lcm {1..nat \real n\}))" by simp also have "\ \ o(\n. c powr real n)" using Lcm_upto.smallo' by (rule landau_o.small.compose) (simp_all add: c filterlim_real_sequentially) also have "(\n. c powr real n) = (\n. c ^ n)" using c \c > 0\ by (subst powr_realpow) auto finally show ?thesis . qed have "(\n. 2 * b * Re (zeta 3) * real (Lcm {1..n}) ^ 3 / 27 ^ n) \ O(\n. real (Lcm {1..n}) ^ 3 / 27 ^ n)" using \b > 0\ Re_zeta_ge_1[of 3] by simp also have "exp 1 < (3 :: real)" using e_approx_32 by (simp add: abs_if split: if_splits) hence "(\n. real (Lcm {1..n}) ^ 3 / 27 ^ n) \ o(\n. (3 ^ n) ^ 3 / 27 ^ n)" unfolding of_nat_power by (intro landau_o.small.divide_right smallo_power Lcm_upto_smallo) auto also have "(\n. (3 ^ n) ^ 3 / 27 ^ n :: real) = (\_. 1)" by (simp add: power_mult [of 3, symmetric] mult.commute[of _ 3] power_mult[of _ 3]) finally have *: "(\n. 2 * b * Re (zeta 3) * real (Lcm {1..n}) ^ 3 / 27 ^ n) \ o(\_. 1)" . have lim: "(\n. 2 * b * Re (zeta 3) * real (Lcm {1..n}) ^ 3 / 27 ^ n) \ 0" using smalloD_tendsto[OF *] by simp moreover have "1 \ real (2 * b) * Re (zeta 3) * real (Lcm {1..n} ^ 3) / 27 ^ n" for n using beukers_key_inequality[of b a] ab \b > 0\ by simp ultimately have "1 \ (0 :: real)" by (intro tendsto_lowerbound[OF lim] always_eventually allI) auto thus False by simp qed end \ No newline at end of file