diff --git a/src/HOL/Complex.thy b/src/HOL/Complex.thy --- a/src/HOL/Complex.thy +++ b/src/HOL/Complex.thy @@ -1,1307 +1,1313 @@ (* Title: HOL/Complex.thy Author: Jacques D. Fleuriot, 2001 University of Edinburgh Author: Lawrence C Paulson, 2003/4 *) section \Complex Numbers: Rectangular and Polar Representations\ theory Complex imports Transcendental begin text \ We use the \<^theory_text>\codatatype\ command to define the type of complex numbers. This allows us to use \<^theory_text>\primcorec\ to define complex functions by defining their real and imaginary result separately. \ codatatype complex = Complex (Re: real) (Im: real) lemma complex_surj: "Complex (Re z) (Im z) = z" by (rule complex.collapse) lemma complex_eqI [intro?]: "Re x = Re y \ Im x = Im y \ x = y" by (rule complex.expand) simp lemma complex_eq_iff: "x = y \ Re x = Re y \ Im x = Im y" by (auto intro: complex.expand) subsection \Addition and Subtraction\ instantiation complex :: ab_group_add begin primcorec zero_complex where "Re 0 = 0" | "Im 0 = 0" primcorec plus_complex where "Re (x + y) = Re x + Re y" | "Im (x + y) = Im x + Im y" primcorec uminus_complex where "Re (- x) = - Re x" | "Im (- x) = - Im x" primcorec minus_complex where "Re (x - y) = Re x - Re y" | "Im (x - y) = Im x - Im y" instance by standard (simp_all add: complex_eq_iff) end subsection \Multiplication and Division\ instantiation complex :: field begin primcorec one_complex where "Re 1 = 1" | "Im 1 = 0" primcorec times_complex where "Re (x * y) = Re x * Re y - Im x * Im y" | "Im (x * y) = Re x * Im y + Im x * Re y" primcorec inverse_complex where "Re (inverse x) = Re x / ((Re x)\<^sup>2 + (Im x)\<^sup>2)" | "Im (inverse x) = - Im x / ((Re x)\<^sup>2 + (Im x)\<^sup>2)" definition "x div y = x * inverse y" for x y :: complex instance by standard (simp_all add: complex_eq_iff divide_complex_def distrib_left distrib_right right_diff_distrib left_diff_distrib power2_eq_square add_divide_distrib [symmetric]) end lemma Re_divide: "Re (x / y) = (Re x * Re y + Im x * Im y) / ((Re y)\<^sup>2 + (Im y)\<^sup>2)" by (simp add: divide_complex_def add_divide_distrib) lemma Im_divide: "Im (x / y) = (Im x * Re y - Re x * Im y) / ((Re y)\<^sup>2 + (Im y)\<^sup>2)" by (simp add: divide_complex_def diff_divide_distrib) lemma Complex_divide: "(x / y) = Complex ((Re x * Re y + Im x * Im y) / ((Re y)\<^sup>2 + (Im y)\<^sup>2)) ((Im x * Re y - Re x * Im y) / ((Re y)\<^sup>2 + (Im y)\<^sup>2))" by (metis Im_divide Re_divide complex_surj) lemma Re_power2: "Re (x ^ 2) = (Re x)^2 - (Im x)^2" by (simp add: power2_eq_square) lemma Im_power2: "Im (x ^ 2) = 2 * Re x * Im x" by (simp add: power2_eq_square) lemma Re_power_real [simp]: "Im x = 0 \ Re (x ^ n) = Re x ^ n " by (induct n) simp_all lemma Im_power_real [simp]: "Im x = 0 \ Im (x ^ n) = 0" by (induct n) simp_all subsection \Scalar Multiplication\ instantiation complex :: real_field begin primcorec scaleR_complex where "Re (scaleR r x) = r * Re x" | "Im (scaleR r x) = r * Im x" instance proof fix a b :: real and x y :: complex show "scaleR a (x + y) = scaleR a x + scaleR a y" by (simp add: complex_eq_iff distrib_left) show "scaleR (a + b) x = scaleR a x + scaleR b x" by (simp add: complex_eq_iff distrib_right) show "scaleR a (scaleR b x) = scaleR (a * b) x" by (simp add: complex_eq_iff mult.assoc) show "scaleR 1 x = x" by (simp add: complex_eq_iff) show "scaleR a x * y = scaleR a (x * y)" by (simp add: complex_eq_iff algebra_simps) show "x * scaleR a y = scaleR a (x * y)" by (simp add: complex_eq_iff algebra_simps) qed end subsection \Numerals, Arithmetic, and Embedding from R\ abbreviation complex_of_real :: "real \ complex" where "complex_of_real \ of_real" declare [[coercion "of_real :: real \ complex"]] declare [[coercion "of_rat :: rat \ complex"]] declare [[coercion "of_int :: int \ complex"]] declare [[coercion "of_nat :: nat \ complex"]] lemma complex_Re_of_nat [simp]: "Re (of_nat n) = of_nat n" by (induct n) simp_all lemma complex_Im_of_nat [simp]: "Im (of_nat n) = 0" by (induct n) simp_all lemma complex_Re_of_int [simp]: "Re (of_int z) = of_int z" by (cases z rule: int_diff_cases) simp lemma complex_Im_of_int [simp]: "Im (of_int z) = 0" by (cases z rule: int_diff_cases) simp lemma complex_Re_numeral [simp]: "Re (numeral v) = numeral v" using complex_Re_of_int [of "numeral v"] by simp lemma complex_Im_numeral [simp]: "Im (numeral v) = 0" using complex_Im_of_int [of "numeral v"] by simp lemma Re_complex_of_real [simp]: "Re (complex_of_real z) = z" by (simp add: of_real_def) lemma Im_complex_of_real [simp]: "Im (complex_of_real z) = 0" by (simp add: of_real_def) lemma Re_divide_numeral [simp]: "Re (z / numeral w) = Re z / numeral w" by (simp add: Re_divide sqr_conv_mult) lemma Im_divide_numeral [simp]: "Im (z / numeral w) = Im z / numeral w" by (simp add: Im_divide sqr_conv_mult) lemma Re_divide_of_nat [simp]: "Re (z / of_nat n) = Re z / of_nat n" by (cases n) (simp_all add: Re_divide field_split_simps power2_eq_square del: of_nat_Suc) lemma Im_divide_of_nat [simp]: "Im (z / of_nat n) = Im z / of_nat n" by (cases n) (simp_all add: Im_divide field_split_simps power2_eq_square del: of_nat_Suc) lemma of_real_Re [simp]: "z \ \ \ of_real (Re z) = z" by (auto simp: Reals_def) lemma complex_Re_fact [simp]: "Re (fact n) = fact n" proof - have "(fact n :: complex) = of_real (fact n)" by simp also have "Re \ = fact n" by (subst Re_complex_of_real) simp_all finally show ?thesis . qed lemma complex_Im_fact [simp]: "Im (fact n) = 0" by (subst of_nat_fact [symmetric]) (simp only: complex_Im_of_nat) lemma Re_prod_Reals: "(\x. x \ A \ f x \ \) \ Re (prod f A) = prod (\x. Re (f x)) A" proof (induction A rule: infinite_finite_induct) case (insert x A) hence "Re (prod f (insert x A)) = Re (f x) * Re (prod f A) - Im (f x) * Im (prod f A)" by simp also from insert.prems have "f x \ \" by simp hence "Im (f x) = 0" by (auto elim!: Reals_cases) also have "Re (prod f A) = (\x\A. Re (f x))" by (intro insert.IH insert.prems) auto finally show ?case using insert.hyps by simp qed auto subsection \The Complex Number $i$\ primcorec imaginary_unit :: complex ("\") where "Re \ = 0" | "Im \ = 1" lemma Complex_eq: "Complex a b = a + \ * b" by (simp add: complex_eq_iff) lemma complex_eq: "a = Re a + \ * Im a" by (simp add: complex_eq_iff) lemma fun_complex_eq: "f = (\x. Re (f x) + \ * Im (f x))" by (simp add: fun_eq_iff complex_eq) lemma i_squared [simp]: "\ * \ = -1" by (simp add: complex_eq_iff) lemma power2_i [simp]: "\\<^sup>2 = -1" by (simp add: power2_eq_square) lemma inverse_i [simp]: "inverse \ = - \" by (rule inverse_unique) simp lemma divide_i [simp]: "x / \ = - \ * x" by (simp add: divide_complex_def) lemma complex_i_mult_minus [simp]: "\ * (\ * x) = - x" by (simp add: mult.assoc [symmetric]) lemma complex_i_not_zero [simp]: "\ \ 0" by (simp add: complex_eq_iff) lemma complex_i_not_one [simp]: "\ \ 1" by (simp add: complex_eq_iff) lemma complex_i_not_numeral [simp]: "\ \ numeral w" by (simp add: complex_eq_iff) lemma complex_i_not_neg_numeral [simp]: "\ \ - numeral w" by (simp add: complex_eq_iff) lemma complex_split_polar: "\r a. z = complex_of_real r * (cos a + \ * sin a)" by (simp add: complex_eq_iff polar_Ex) lemma i_even_power [simp]: "\ ^ (n * 2) = (-1) ^ n" by (metis mult.commute power2_i power_mult) lemma Re_i_times [simp]: "Re (\ * z) = - Im z" by simp lemma Im_i_times [simp]: "Im (\ * z) = Re z" by simp lemma i_times_eq_iff: "\ * w = z \ w = - (\ * z)" by auto lemma divide_numeral_i [simp]: "z / (numeral n * \) = - (\ * z) / numeral n" by (metis divide_divide_eq_left divide_i mult.commute mult_minus_right) lemma imaginary_eq_real_iff [simp]: assumes "y \ Reals" "x \ Reals" shows "\ * y = x \ x=0 \ y=0" using assms unfolding Reals_def apply clarify apply (rule iffI) apply (metis Im_complex_of_real Im_i_times Re_complex_of_real mult_eq_0_iff of_real_0) by simp lemma real_eq_imaginary_iff [simp]: assumes "y \ Reals" "x \ Reals" shows "x = \ * y \ x=0 \ y=0" using assms imaginary_eq_real_iff by fastforce subsection \Vector Norm\ instantiation complex :: real_normed_field begin definition "norm z = sqrt ((Re z)\<^sup>2 + (Im z)\<^sup>2)" abbreviation cmod :: "complex \ real" where "cmod \ norm" definition complex_sgn_def: "sgn x = x /\<^sub>R cmod x" definition dist_complex_def: "dist x y = cmod (x - y)" definition uniformity_complex_def [code del]: "(uniformity :: (complex \ complex) filter) = (INF e\{0 <..}. principal {(x, y). dist x y < e})" definition open_complex_def [code del]: "open (U :: complex set) \ (\x\U. eventually (\(x', y). x' = x \ y \ U) uniformity)" instance proof fix r :: real and x y :: complex and S :: "complex set" show "(norm x = 0) = (x = 0)" by (simp add: norm_complex_def complex_eq_iff) show "norm (x + y) \ norm x + norm y" by (simp add: norm_complex_def complex_eq_iff real_sqrt_sum_squares_triangle_ineq) show "norm (scaleR r x) = \r\ * norm x" by (simp add: norm_complex_def complex_eq_iff power_mult_distrib distrib_left [symmetric] real_sqrt_mult) show "norm (x * y) = norm x * norm y" by (simp add: norm_complex_def complex_eq_iff real_sqrt_mult [symmetric] power2_eq_square algebra_simps) qed (rule complex_sgn_def dist_complex_def open_complex_def uniformity_complex_def)+ end declare uniformity_Abort[where 'a = complex, code] lemma norm_ii [simp]: "norm \ = 1" by (simp add: norm_complex_def) lemma cmod_unit_one: "cmod (cos a + \ * sin a) = 1" by (simp add: norm_complex_def) lemma cmod_complex_polar: "cmod (r * (cos a + \ * sin a)) = \r\" by (simp add: norm_mult cmod_unit_one) lemma complex_Re_le_cmod: "Re x \ cmod x" unfolding norm_complex_def by (rule real_sqrt_sum_squares_ge1) lemma complex_mod_minus_le_complex_mod: "- cmod x \ cmod x" by (rule order_trans [OF _ norm_ge_zero]) simp lemma complex_mod_triangle_ineq2: "cmod (b + a) - cmod b \ cmod a" by (rule ord_le_eq_trans [OF norm_triangle_ineq2]) simp lemma abs_Re_le_cmod: "\Re x\ \ cmod x" by (simp add: norm_complex_def) lemma abs_Im_le_cmod: "\Im x\ \ cmod x" by (simp add: norm_complex_def) lemma cmod_le: "cmod z \ \Re z\ + \Im z\" apply (subst complex_eq) apply (rule order_trans) apply (rule norm_triangle_ineq) apply (simp add: norm_mult) done lemma cmod_eq_Re: "Im z = 0 \ cmod z = \Re z\" by (simp add: norm_complex_def) lemma cmod_eq_Im: "Re z = 0 \ cmod z = \Im z\" by (simp add: norm_complex_def) lemma cmod_power2: "(cmod z)\<^sup>2 = (Re z)\<^sup>2 + (Im z)\<^sup>2" by (simp add: norm_complex_def) lemma cmod_plus_Re_le_0_iff: "cmod z + Re z \ 0 \ Re z = - cmod z" using abs_Re_le_cmod[of z] by auto lemma cmod_Re_le_iff: "Im x = Im y \ cmod x \ cmod y \ \Re x\ \ \Re y\" by (metis add.commute add_le_cancel_left norm_complex_def real_sqrt_abs real_sqrt_le_iff) lemma cmod_Im_le_iff: "Re x = Re y \ cmod x \ cmod y \ \Im x\ \ \Im y\" by (metis add_le_cancel_left norm_complex_def real_sqrt_abs real_sqrt_le_iff) lemma Im_eq_0: "\Re z\ = cmod z \ Im z = 0" by (subst (asm) power_eq_iff_eq_base[symmetric, where n=2]) (auto simp add: norm_complex_def) lemma abs_sqrt_wlog: "(\x. x \ 0 \ P x (x\<^sup>2)) \ P \x\ (x\<^sup>2)" for x::"'a::linordered_idom" by (metis abs_ge_zero power2_abs) lemma complex_abs_le_norm: "\Re z\ + \Im z\ \ sqrt 2 * norm z" unfolding norm_complex_def apply (rule abs_sqrt_wlog [where x="Re z"]) apply (rule abs_sqrt_wlog [where x="Im z"]) apply (rule power2_le_imp_le) apply (simp_all add: power2_sum add.commute sum_squares_bound real_sqrt_mult [symmetric]) done lemma complex_unit_circle: "z \ 0 \ (Re z / cmod z)\<^sup>2 + (Im z / cmod z)\<^sup>2 = 1" by (simp add: norm_complex_def complex_eq_iff power2_eq_square add_divide_distrib [symmetric]) text \Properties of complex signum.\ lemma sgn_eq: "sgn z = z / complex_of_real (cmod z)" by (simp add: sgn_div_norm divide_inverse scaleR_conv_of_real mult.commute) lemma Re_sgn [simp]: "Re(sgn z) = Re(z)/cmod z" by (simp add: complex_sgn_def divide_inverse) lemma Im_sgn [simp]: "Im(sgn z) = Im(z)/cmod z" by (simp add: complex_sgn_def divide_inverse) subsection \Absolute value\ instantiation complex :: field_abs_sgn begin definition abs_complex :: "complex \ complex" where "abs_complex = of_real \ norm" instance apply standard apply (auto simp add: abs_complex_def complex_sgn_def norm_mult) apply (auto simp add: scaleR_conv_of_real field_simps) done end subsection \Completeness of the Complexes\ lemma bounded_linear_Re: "bounded_linear Re" by (rule bounded_linear_intro [where K=1]) (simp_all add: norm_complex_def) lemma bounded_linear_Im: "bounded_linear Im" by (rule bounded_linear_intro [where K=1]) (simp_all add: norm_complex_def) lemmas Cauchy_Re = bounded_linear.Cauchy [OF bounded_linear_Re] lemmas Cauchy_Im = bounded_linear.Cauchy [OF bounded_linear_Im] lemmas tendsto_Re [tendsto_intros] = bounded_linear.tendsto [OF bounded_linear_Re] lemmas tendsto_Im [tendsto_intros] = bounded_linear.tendsto [OF bounded_linear_Im] lemmas isCont_Re [simp] = bounded_linear.isCont [OF bounded_linear_Re] lemmas isCont_Im [simp] = bounded_linear.isCont [OF bounded_linear_Im] lemmas continuous_Re [simp] = bounded_linear.continuous [OF bounded_linear_Re] lemmas continuous_Im [simp] = bounded_linear.continuous [OF bounded_linear_Im] lemmas continuous_on_Re [continuous_intros] = bounded_linear.continuous_on[OF bounded_linear_Re] lemmas continuous_on_Im [continuous_intros] = bounded_linear.continuous_on[OF bounded_linear_Im] lemmas has_derivative_Re [derivative_intros] = bounded_linear.has_derivative[OF bounded_linear_Re] lemmas has_derivative_Im [derivative_intros] = bounded_linear.has_derivative[OF bounded_linear_Im] lemmas sums_Re = bounded_linear.sums [OF bounded_linear_Re] lemmas sums_Im = bounded_linear.sums [OF bounded_linear_Im] lemma tendsto_Complex [tendsto_intros]: "(f \ a) F \ (g \ b) F \ ((\x. Complex (f x) (g x)) \ Complex a b) F" unfolding Complex_eq by (auto intro!: tendsto_intros) lemma tendsto_complex_iff: "(f \ x) F \ (((\x. Re (f x)) \ Re x) F \ ((\x. Im (f x)) \ Im x) F)" proof safe assume "((\x. Re (f x)) \ Re x) F" "((\x. Im (f x)) \ Im x) F" from tendsto_Complex[OF this] show "(f \ x) F" unfolding complex.collapse . qed (auto intro: tendsto_intros) lemma continuous_complex_iff: "continuous F f \ continuous F (\x. Re (f x)) \ continuous F (\x. Im (f x))" by (simp only: continuous_def tendsto_complex_iff) lemma continuous_on_of_real_o_iff [simp]: "continuous_on S (\x. complex_of_real (g x)) = continuous_on S g" using continuous_on_Re continuous_on_of_real by fastforce lemma continuous_on_of_real_id [simp]: "continuous_on S (of_real :: real \ 'a::real_normed_algebra_1)" by (rule continuous_on_of_real [OF continuous_on_id]) lemma has_vector_derivative_complex_iff: "(f has_vector_derivative x) F \ ((\x. Re (f x)) has_field_derivative (Re x)) F \ ((\x. Im (f x)) has_field_derivative (Im x)) F" by (simp add: has_vector_derivative_def has_field_derivative_def has_derivative_def tendsto_complex_iff algebra_simps bounded_linear_scaleR_left bounded_linear_mult_right) lemma has_field_derivative_Re[derivative_intros]: "(f has_vector_derivative D) F \ ((\x. Re (f x)) has_field_derivative (Re D)) F" unfolding has_vector_derivative_complex_iff by safe lemma has_field_derivative_Im[derivative_intros]: "(f has_vector_derivative D) F \ ((\x. Im (f x)) has_field_derivative (Im D)) F" unfolding has_vector_derivative_complex_iff by safe instance complex :: banach proof fix X :: "nat \ complex" assume X: "Cauchy X" then have "(\n. Complex (Re (X n)) (Im (X n))) \ Complex (lim (\n. Re (X n))) (lim (\n. Im (X n)))" by (intro tendsto_Complex convergent_LIMSEQ_iff[THEN iffD1] Cauchy_convergent_iff[THEN iffD1] Cauchy_Re Cauchy_Im) then show "convergent X" unfolding complex.collapse by (rule convergentI) qed declare DERIV_power[where 'a=complex, unfolded of_nat_def[symmetric], derivative_intros] subsection \Complex Conjugation\ primcorec cnj :: "complex \ complex" where "Re (cnj z) = Re z" | "Im (cnj z) = - Im z" lemma complex_cnj_cancel_iff [simp]: "cnj x = cnj y \ x = y" by (simp add: complex_eq_iff) lemma complex_cnj_cnj [simp]: "cnj (cnj z) = z" by (simp add: complex_eq_iff) lemma complex_cnj_zero [simp]: "cnj 0 = 0" by (simp add: complex_eq_iff) lemma complex_cnj_zero_iff [iff]: "cnj z = 0 \ z = 0" by (simp add: complex_eq_iff) lemma complex_cnj_one_iff [simp]: "cnj z = 1 \ z = 1" by (simp add: complex_eq_iff) lemma complex_cnj_add [simp]: "cnj (x + y) = cnj x + cnj y" by (simp add: complex_eq_iff) lemma cnj_sum [simp]: "cnj (sum f s) = (\x\s. cnj (f x))" by (induct s rule: infinite_finite_induct) auto lemma complex_cnj_diff [simp]: "cnj (x - y) = cnj x - cnj y" by (simp add: complex_eq_iff) lemma complex_cnj_minus [simp]: "cnj (- x) = - cnj x" by (simp add: complex_eq_iff) lemma complex_cnj_one [simp]: "cnj 1 = 1" by (simp add: complex_eq_iff) lemma complex_cnj_mult [simp]: "cnj (x * y) = cnj x * cnj y" by (simp add: complex_eq_iff) lemma cnj_prod [simp]: "cnj (prod f s) = (\x\s. cnj (f x))" by (induct s rule: infinite_finite_induct) auto lemma complex_cnj_inverse [simp]: "cnj (inverse x) = inverse (cnj x)" by (simp add: complex_eq_iff) lemma complex_cnj_divide [simp]: "cnj (x / y) = cnj x / cnj y" by (simp add: divide_complex_def) lemma complex_cnj_power [simp]: "cnj (x ^ n) = cnj x ^ n" by (induct n) simp_all lemma complex_cnj_of_nat [simp]: "cnj (of_nat n) = of_nat n" by (simp add: complex_eq_iff) lemma complex_cnj_of_int [simp]: "cnj (of_int z) = of_int z" by (simp add: complex_eq_iff) lemma complex_cnj_numeral [simp]: "cnj (numeral w) = numeral w" by (simp add: complex_eq_iff) lemma complex_cnj_neg_numeral [simp]: "cnj (- numeral w) = - numeral w" by (simp add: complex_eq_iff) lemma complex_cnj_scaleR [simp]: "cnj (scaleR r x) = scaleR r (cnj x)" by (simp add: complex_eq_iff) lemma complex_mod_cnj [simp]: "cmod (cnj z) = cmod z" by (simp add: norm_complex_def) lemma complex_cnj_complex_of_real [simp]: "cnj (of_real x) = of_real x" by (simp add: complex_eq_iff) lemma complex_cnj_i [simp]: "cnj \ = - \" by (simp add: complex_eq_iff) lemma complex_add_cnj: "z + cnj z = complex_of_real (2 * Re z)" by (simp add: complex_eq_iff) lemma complex_diff_cnj: "z - cnj z = complex_of_real (2 * Im z) * \" by (simp add: complex_eq_iff) +lemma Ints_cnj [intro]: "x \ \ \ cnj x \ \" + by (auto elim!: Ints_cases) + +lemma cnj_in_Ints_iff [simp]: "cnj x \ \ \ x \ \" + using Ints_cnj[of x] Ints_cnj[of "cnj x"] by auto + lemma complex_mult_cnj: "z * cnj z = complex_of_real ((Re z)\<^sup>2 + (Im z)\<^sup>2)" by (simp add: complex_eq_iff power2_eq_square) lemma cnj_add_mult_eq_Re: "z * cnj w + cnj z * w = 2 * Re (z * cnj w)" by (rule complex_eqI) auto lemma complex_mod_mult_cnj: "cmod (z * cnj z) = (cmod z)\<^sup>2" by (simp add: norm_mult power2_eq_square) lemma complex_mod_sqrt_Re_mult_cnj: "cmod z = sqrt (Re (z * cnj z))" by (simp add: norm_complex_def power2_eq_square) lemma complex_In_mult_cnj_zero [simp]: "Im (z * cnj z) = 0" by simp lemma complex_cnj_fact [simp]: "cnj (fact n) = fact n" by (subst of_nat_fact [symmetric], subst complex_cnj_of_nat) simp lemma complex_cnj_pochhammer [simp]: "cnj (pochhammer z n) = pochhammer (cnj z) n" by (induct n arbitrary: z) (simp_all add: pochhammer_rec) lemma bounded_linear_cnj: "bounded_linear cnj" using complex_cnj_add complex_cnj_scaleR by (rule bounded_linear_intro [where K=1]) simp lemma linear_cnj: "linear cnj" using bounded_linear.linear[OF bounded_linear_cnj] . lemmas tendsto_cnj [tendsto_intros] = bounded_linear.tendsto [OF bounded_linear_cnj] and isCont_cnj [simp] = bounded_linear.isCont [OF bounded_linear_cnj] and continuous_cnj [simp, continuous_intros] = bounded_linear.continuous [OF bounded_linear_cnj] and continuous_on_cnj [simp, continuous_intros] = bounded_linear.continuous_on [OF bounded_linear_cnj] and has_derivative_cnj [simp, derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_cnj] lemma lim_cnj: "((\x. cnj(f x)) \ cnj l) F \ (f \ l) F" by (simp add: tendsto_iff dist_complex_def complex_cnj_diff [symmetric] del: complex_cnj_diff) lemma sums_cnj: "((\x. cnj(f x)) sums cnj l) \ (f sums l)" by (simp add: sums_def lim_cnj cnj_sum [symmetric] del: cnj_sum) lemma differentiable_cnj_iff: "(\z. cnj (f z)) differentiable at x within A \ f differentiable at x within A" proof assume "(\z. cnj (f z)) differentiable at x within A" then obtain D where "((\z. cnj (f z)) has_derivative D) (at x within A)" by (auto simp: differentiable_def) from has_derivative_cnj[OF this] show "f differentiable at x within A" by (auto simp: differentiable_def) next assume "f differentiable at x within A" then obtain D where "(f has_derivative D) (at x within A)" by (auto simp: differentiable_def) from has_derivative_cnj[OF this] show "(\z. cnj (f z)) differentiable at x within A" by (auto simp: differentiable_def) qed lemma has_vector_derivative_cnj [derivative_intros]: assumes "(f has_vector_derivative f') (at z within A)" shows "((\z. cnj (f z)) has_vector_derivative cnj f') (at z within A)" using assms by (auto simp: has_vector_derivative_complex_iff intro: derivative_intros) subsection \Basic Lemmas\ lemma complex_eq_0: "z=0 \ (Re z)\<^sup>2 + (Im z)\<^sup>2 = 0" by (metis zero_complex.sel complex_eqI sum_power2_eq_zero_iff) lemma complex_neq_0: "z\0 \ (Re z)\<^sup>2 + (Im z)\<^sup>2 > 0" by (metis complex_eq_0 less_numeral_extra(3) sum_power2_gt_zero_iff) lemma complex_norm_square: "of_real ((norm z)\<^sup>2) = z * cnj z" by (cases z) (auto simp: complex_eq_iff norm_complex_def power2_eq_square[symmetric] of_real_power[symmetric] simp del: of_real_power) lemma complex_div_cnj: "a / b = (a * cnj b) / (norm b)\<^sup>2" using complex_norm_square by auto lemma Re_complex_div_eq_0: "Re (a / b) = 0 \ Re (a * cnj b) = 0" by (auto simp add: Re_divide) lemma Im_complex_div_eq_0: "Im (a / b) = 0 \ Im (a * cnj b) = 0" by (auto simp add: Im_divide) lemma complex_div_gt_0: "(Re (a / b) > 0 \ Re (a * cnj b) > 0) \ (Im (a / b) > 0 \ Im (a * cnj b) > 0)" proof (cases "b = 0") case True then show ?thesis by auto next case False then have "0 < (Re b)\<^sup>2 + (Im b)\<^sup>2" by (simp add: complex_eq_iff sum_power2_gt_zero_iff) then show ?thesis by (simp add: Re_divide Im_divide zero_less_divide_iff) qed lemma Re_complex_div_gt_0: "Re (a / b) > 0 \ Re (a * cnj b) > 0" and Im_complex_div_gt_0: "Im (a / b) > 0 \ Im (a * cnj b) > 0" using complex_div_gt_0 by auto lemma Re_complex_div_ge_0: "Re (a / b) \ 0 \ Re (a * cnj b) \ 0" by (metis le_less Re_complex_div_eq_0 Re_complex_div_gt_0) lemma Im_complex_div_ge_0: "Im (a / b) \ 0 \ Im (a * cnj b) \ 0" by (metis Im_complex_div_eq_0 Im_complex_div_gt_0 le_less) lemma Re_complex_div_lt_0: "Re (a / b) < 0 \ Re (a * cnj b) < 0" by (metis less_asym neq_iff Re_complex_div_eq_0 Re_complex_div_gt_0) lemma Im_complex_div_lt_0: "Im (a / b) < 0 \ Im (a * cnj b) < 0" by (metis Im_complex_div_eq_0 Im_complex_div_gt_0 less_asym neq_iff) lemma Re_complex_div_le_0: "Re (a / b) \ 0 \ Re (a * cnj b) \ 0" by (metis not_le Re_complex_div_gt_0) lemma Im_complex_div_le_0: "Im (a / b) \ 0 \ Im (a * cnj b) \ 0" by (metis Im_complex_div_gt_0 not_le) lemma Re_divide_of_real [simp]: "Re (z / of_real r) = Re z / r" by (simp add: Re_divide power2_eq_square) lemma Im_divide_of_real [simp]: "Im (z / of_real r) = Im z / r" by (simp add: Im_divide power2_eq_square) lemma Re_divide_Reals [simp]: "r \ \ \ Re (z / r) = Re z / Re r" by (metis Re_divide_of_real of_real_Re) lemma Im_divide_Reals [simp]: "r \ \ \ Im (z / r) = Im z / Re r" by (metis Im_divide_of_real of_real_Re) lemma Re_sum[simp]: "Re (sum f s) = (\x\s. Re (f x))" by (induct s rule: infinite_finite_induct) auto lemma Im_sum[simp]: "Im (sum f s) = (\x\s. Im(f x))" by (induct s rule: infinite_finite_induct) auto lemma sums_complex_iff: "f sums x \ ((\x. Re (f x)) sums Re x) \ ((\x. Im (f x)) sums Im x)" unfolding sums_def tendsto_complex_iff Im_sum Re_sum .. lemma summable_complex_iff: "summable f \ summable (\x. Re (f x)) \ summable (\x. Im (f x))" unfolding summable_def sums_complex_iff[abs_def] by (metis complex.sel) lemma summable_complex_of_real [simp]: "summable (\n. complex_of_real (f n)) \ summable f" unfolding summable_complex_iff by simp lemma summable_Re: "summable f \ summable (\x. Re (f x))" unfolding summable_complex_iff by blast lemma summable_Im: "summable f \ summable (\x. Im (f x))" unfolding summable_complex_iff by blast lemma complex_is_Nat_iff: "z \ \ \ Im z = 0 \ (\i. Re z = of_nat i)" by (auto simp: Nats_def complex_eq_iff) lemma complex_is_Int_iff: "z \ \ \ Im z = 0 \ (\i. Re z = of_int i)" by (auto simp: Ints_def complex_eq_iff) lemma complex_is_Real_iff: "z \ \ \ Im z = 0" by (auto simp: Reals_def complex_eq_iff) lemma Reals_cnj_iff: "z \ \ \ cnj z = z" by (auto simp: complex_is_Real_iff complex_eq_iff) lemma in_Reals_norm: "z \ \ \ norm z = \Re z\" by (simp add: complex_is_Real_iff norm_complex_def) lemma Re_Reals_divide: "r \ \ \ Re (r / z) = Re r * Re z / (norm z)\<^sup>2" by (simp add: Re_divide complex_is_Real_iff cmod_power2) lemma Im_Reals_divide: "r \ \ \ Im (r / z) = -Re r * Im z / (norm z)\<^sup>2" by (simp add: Im_divide complex_is_Real_iff cmod_power2) lemma series_comparison_complex: fixes f:: "nat \ 'a::banach" assumes sg: "summable g" and "\n. g n \ \" "\n. Re (g n) \ 0" and fg: "\n. n \ N \ norm(f n) \ norm(g n)" shows "summable f" proof - have g: "\n. cmod (g n) = Re (g n)" using assms by (metis abs_of_nonneg in_Reals_norm) show ?thesis apply (rule summable_comparison_test' [where g = "\n. norm (g n)" and N=N]) using sg apply (auto simp: summable_def) apply (rule_tac x = "Re s" in exI) apply (auto simp: g sums_Re) apply (metis fg g) done qed subsection \Polar Form for Complex Numbers\ lemma complex_unimodular_polar: assumes "norm z = 1" obtains t where "0 \ t" "t < 2 * pi" "z = Complex (cos t) (sin t)" by (metis cmod_power2 one_power2 complex_surj sincos_total_2pi [of "Re z" "Im z"] assms) subsubsection \$\cos \theta + i \sin \theta$\ primcorec cis :: "real \ complex" where "Re (cis a) = cos a" | "Im (cis a) = sin a" lemma cis_zero [simp]: "cis 0 = 1" by (simp add: complex_eq_iff) lemma norm_cis [simp]: "norm (cis a) = 1" by (simp add: norm_complex_def) lemma sgn_cis [simp]: "sgn (cis a) = cis a" by (simp add: sgn_div_norm) lemma cis_2pi [simp]: "cis (2 * pi) = 1" by (simp add: cis.ctr complex_eq_iff) lemma cis_neq_zero [simp]: "cis a \ 0" by (metis norm_cis norm_zero zero_neq_one) lemma cis_cnj: "cnj (cis t) = cis (-t)" by (simp add: complex_eq_iff) lemma cis_mult: "cis a * cis b = cis (a + b)" by (simp add: complex_eq_iff cos_add sin_add) lemma DeMoivre: "(cis a) ^ n = cis (real n * a)" by (induct n) (simp_all add: algebra_simps cis_mult) lemma cis_inverse [simp]: "inverse (cis a) = cis (- a)" by (simp add: complex_eq_iff) lemma cis_divide: "cis a / cis b = cis (a - b)" by (simp add: divide_complex_def cis_mult) lemma cos_n_Re_cis_pow_n: "cos (real n * a) = Re (cis a ^ n)" by (auto simp add: DeMoivre) lemma sin_n_Im_cis_pow_n: "sin (real n * a) = Im (cis a ^ n)" by (auto simp add: DeMoivre) lemma cis_pi [simp]: "cis pi = -1" by (simp add: complex_eq_iff) lemma cis_pi_half[simp]: "cis (pi / 2) = \" by (simp add: cis.ctr complex_eq_iff) lemma cis_minus_pi_half[simp]: "cis (-(pi / 2)) = -\" by (simp add: cis.ctr complex_eq_iff) lemma cis_multiple_2pi[simp]: "n \ \ \ cis (2 * pi * n) = 1" by (auto elim!: Ints_cases simp: cis.ctr one_complex.ctr) subsubsection \$r(\cos \theta + i \sin \theta)$\ definition rcis :: "real \ real \ complex" where "rcis r a = complex_of_real r * cis a" lemma Re_rcis [simp]: "Re(rcis r a) = r * cos a" by (simp add: rcis_def) lemma Im_rcis [simp]: "Im(rcis r a) = r * sin a" by (simp add: rcis_def) lemma rcis_Ex: "\r a. z = rcis r a" by (simp add: complex_eq_iff polar_Ex) lemma complex_mod_rcis [simp]: "cmod (rcis r a) = \r\" by (simp add: rcis_def norm_mult) lemma cis_rcis_eq: "cis a = rcis 1 a" by (simp add: rcis_def) lemma rcis_mult: "rcis r1 a * rcis r2 b = rcis (r1 * r2) (a + b)" by (simp add: rcis_def cis_mult) lemma rcis_zero_mod [simp]: "rcis 0 a = 0" by (simp add: rcis_def) lemma rcis_zero_arg [simp]: "rcis r 0 = complex_of_real r" by (simp add: rcis_def) lemma rcis_eq_zero_iff [simp]: "rcis r a = 0 \ r = 0" by (simp add: rcis_def) lemma DeMoivre2: "(rcis r a) ^ n = rcis (r ^ n) (real n * a)" by (simp add: rcis_def power_mult_distrib DeMoivre) lemma rcis_inverse: "inverse(rcis r a) = rcis (1 / r) (- a)" by (simp add: divide_inverse rcis_def) lemma rcis_divide: "rcis r1 a / rcis r2 b = rcis (r1 / r2) (a - b)" by (simp add: rcis_def cis_divide [symmetric]) subsubsection \Complex exponential\ lemma exp_Reals_eq: assumes "z \ \" shows "exp z = of_real (exp (Re z))" using assms by (auto elim!: Reals_cases simp: exp_of_real) lemma cis_conv_exp: "cis b = exp (\ * b)" proof - have "(\ * complex_of_real b) ^ n /\<^sub>R fact n = of_real (cos_coeff n * b^n) + \ * of_real (sin_coeff n * b^n)" for n :: nat proof - have "\ ^ n = fact n *\<^sub>R (cos_coeff n + \ * sin_coeff n)" by (induct n) (simp_all add: sin_coeff_Suc cos_coeff_Suc complex_eq_iff Re_divide Im_divide field_simps power2_eq_square add_nonneg_eq_0_iff) then show ?thesis by (simp add: field_simps) qed then show ?thesis using sin_converges [of b] cos_converges [of b] by (auto simp add: Complex_eq cis.ctr exp_def simp del: of_real_mult intro!: sums_unique sums_add sums_mult sums_of_real) qed lemma exp_eq_polar: "exp z = exp (Re z) * cis (Im z)" unfolding cis_conv_exp exp_of_real [symmetric] mult_exp_exp by (cases z) (simp add: Complex_eq) lemma Re_exp: "Re (exp z) = exp (Re z) * cos (Im z)" unfolding exp_eq_polar by simp lemma Im_exp: "Im (exp z) = exp (Re z) * sin (Im z)" unfolding exp_eq_polar by simp lemma norm_cos_sin [simp]: "norm (Complex (cos t) (sin t)) = 1" by (simp add: norm_complex_def) lemma norm_exp_eq_Re [simp]: "norm (exp z) = exp (Re z)" by (simp add: cis.code cmod_complex_polar exp_eq_polar Complex_eq) lemma complex_exp_exists: "\a r. z = complex_of_real r * exp a" apply (insert rcis_Ex [of z]) apply (auto simp add: exp_eq_polar rcis_def mult.assoc [symmetric]) apply (rule_tac x = "\ * complex_of_real a" in exI) apply auto done lemma exp_pi_i [simp]: "exp (of_real pi * \) = -1" by (metis cis_conv_exp cis_pi mult.commute) lemma exp_pi_i' [simp]: "exp (\ * of_real pi) = -1" using cis_conv_exp cis_pi by auto lemma exp_two_pi_i [simp]: "exp (2 * of_real pi * \) = 1" by (simp add: exp_eq_polar complex_eq_iff) lemma exp_two_pi_i' [simp]: "exp (\ * (of_real pi * 2)) = 1" by (metis exp_two_pi_i mult.commute) lemma continuous_on_cis [continuous_intros]: "continuous_on A f \ continuous_on A (\x. cis (f x))" by (auto simp: cis_conv_exp intro!: continuous_intros) subsubsection \Complex argument\ definition arg :: "complex \ real" where "arg z = (if z = 0 then 0 else (SOME a. sgn z = cis a \ - pi < a \ a \ pi))" lemma arg_zero: "arg 0 = 0" by (simp add: arg_def) lemma arg_unique: assumes "sgn z = cis x" and "-pi < x" and "x \ pi" shows "arg z = x" proof - from assms have "z \ 0" by auto have "(SOME a. sgn z = cis a \ -pi < a \ a \ pi) = x" proof fix a define d where "d = a - x" assume a: "sgn z = cis a \ - pi < a \ a \ pi" from a assms have "- (2*pi) < d \ d < 2*pi" unfolding d_def by simp moreover from a assms have "cos a = cos x" and "sin a = sin x" by (simp_all add: complex_eq_iff) then have cos: "cos d = 1" by (simp add: d_def cos_diff) moreover from cos have "sin d = 0" by (rule cos_one_sin_zero) ultimately have "d = 0" by (auto simp: sin_zero_iff elim!: evenE dest!: less_2_cases) then show "a = x" by (simp add: d_def) qed (simp add: assms del: Re_sgn Im_sgn) with \z \ 0\ show "arg z = x" by (simp add: arg_def) qed lemma arg_correct: assumes "z \ 0" shows "sgn z = cis (arg z) \ -pi < arg z \ arg z \ pi" proof (simp add: arg_def assms, rule someI_ex) obtain r a where z: "z = rcis r a" using rcis_Ex by fast with assms have "r \ 0" by auto define b where "b = (if 0 < r then a else a + pi)" have b: "sgn z = cis b" using \r \ 0\ by (simp add: z b_def rcis_def of_real_def sgn_scaleR sgn_if complex_eq_iff) have cis_2pi_nat: "cis (2 * pi * real_of_nat n) = 1" for n by (induct n) (simp_all add: distrib_left cis_mult [symmetric] complex_eq_iff) have cis_2pi_int: "cis (2 * pi * real_of_int x) = 1" for x by (cases x rule: int_diff_cases) (simp add: right_diff_distrib cis_divide [symmetric] cis_2pi_nat) define c where "c = b - 2 * pi * of_int \(b - pi) / (2 * pi)\" have "sgn z = cis c" by (simp add: b c_def cis_divide [symmetric] cis_2pi_int) moreover have "- pi < c \ c \ pi" using ceiling_correct [of "(b - pi) / (2*pi)"] by (simp add: c_def less_divide_eq divide_le_eq algebra_simps del: le_of_int_ceiling) ultimately show "\a. sgn z = cis a \ -pi < a \ a \ pi" by fast qed lemma arg_bounded: "- pi < arg z \ arg z \ pi" by (cases "z = 0") (simp_all add: arg_zero arg_correct) lemma cis_arg: "z \ 0 \ cis (arg z) = sgn z" by (simp add: arg_correct) lemma rcis_cmod_arg: "rcis (cmod z) (arg z) = z" by (cases "z = 0") (simp_all add: rcis_def cis_arg sgn_div_norm of_real_def) lemma cos_arg_i_mult_zero [simp]: "y \ 0 \ Re y = 0 \ cos (arg y) = 0" using cis_arg [of y] by (simp add: complex_eq_iff) subsection \Complex n-th roots\ lemma bij_betw_roots_unity: assumes "n > 0" shows "bij_betw (\k. cis (2 * pi * real k / real n)) {.. = cis (2*pi*(real k - real l)/n)" using assms by (simp add: field_simps cis_divide) finally have "cos (2*pi*(real k - real l) / n) = 1" by (simp add: complex_eq_iff) then obtain m :: int where "2 * pi * (real k - real l) / real n = real_of_int m * 2 * pi" by (subst (asm) cos_one_2pi_int) blast hence "real_of_int (int k - int l) = real_of_int (m * int n)" unfolding of_int_diff of_int_mult using assms by (simp add: nonzero_divide_eq_eq) also note of_int_eq_iff finally have *: "abs m * n = abs (int k - int l)" by (simp add: abs_mult) also have "\ < int n" using kl by linarith finally have "m = 0" using assms by simp with * show "k = l" by simp qed have subset: "?f ` {.. {z. z ^ n = 1}" proof safe fix k :: nat have "cis (2 * pi * real k / real n) ^ n = cis (2 * pi) ^ k" using assms by (simp add: DeMoivre mult_ac) also have "cis (2 * pi) = 1" by (simp add: complex_eq_iff) finally show "?f k ^ n = 1" by simp qed have "n = card {.. \ card {z::complex. z ^ n = 1}" by (intro card_inj_on_le[OF inj]) (auto simp: finite_roots_unity) finally have card: "card {z::complex. z ^ n = 1} = n" using assms by (intro antisym card_roots_unity) auto have "card (?f ` {.. 0" shows "card {z::complex. z ^ n = 1} = n" using bij_betw_same_card [OF bij_betw_roots_unity [OF assms]] by simp lemma bij_betw_nth_root_unity: fixes c :: complex and n :: nat assumes c: "c \ 0" and n: "n > 0" defines "c' \ root n (norm c) * cis (arg c / n)" shows "bij_betw (\z. c' * z) {z. z ^ n = 1} {z. z ^ n = c}" proof - have "c' ^ n = of_real (root n (norm c) ^ n) * cis (arg c)" unfolding of_real_power using n by (simp add: c'_def power_mult_distrib DeMoivre) also from n have "root n (norm c) ^ n = norm c" by simp also from c have "of_real \ * cis (arg c) = c" by (simp add: cis_arg Complex.sgn_eq) finally have [simp]: "c' ^ n = c" . show ?thesis unfolding bij_betw_def inj_on_def proof safe fix z :: complex assume "z ^ n = 1" hence "(c' * z) ^ n = c' ^ n" by (simp add: power_mult_distrib) also have "c' ^ n = of_real (root n (norm c) ^ n) * cis (arg c)" unfolding of_real_power using n by (simp add: c'_def power_mult_distrib DeMoivre) also from n have "root n (norm c) ^ n = norm c" by simp also from c have "\ * cis (arg c) = c" by (simp add: cis_arg Complex.sgn_eq) finally show "(c' * z) ^ n = c" . next fix z assume z: "c = z ^ n" define z' where "z' = z / c'" from c and n have "c' \ 0" by (auto simp: c'_def) with n c have "z = c' * z'" and "z' ^ n = 1" by (auto simp: z'_def power_divide z) thus "z \ (\z. c' * z) ` {z. z ^ n = 1}" by blast qed (insert c n, auto simp: c'_def) qed lemma finite_nth_roots [intro]: assumes "n > 0" shows "finite {z::complex. z ^ n = c}" proof (cases "c = 0") case True with assms have "{z::complex. z ^ n = c} = {0}" by auto thus ?thesis by simp next case False from assms have "finite {z::complex. z ^ n = 1}" by (intro finite_roots_unity) simp_all also have "?this \ ?thesis" by (rule bij_betw_finite, rule bij_betw_nth_root_unity) fact+ finally show ?thesis . qed lemma card_nth_roots: assumes "c \ 0" "n > 0" shows "card {z::complex. z ^ n = c} = n" proof - have "card {z. z ^ n = c} = card {z::complex. z ^ n = 1}" by (rule sym, rule bij_betw_same_card, rule bij_betw_nth_root_unity) fact+ also have "\ = n" by (rule card_roots_unity_eq) fact+ finally show ?thesis . qed lemma sum_roots_unity: assumes "n > 1" shows "\{z::complex. z ^ n = 1} = 0" proof - define \ where "\ = cis (2 * pi / real n)" have [simp]: "\ \ 1" proof assume "\ = 1" with assms obtain k :: int where "2 * pi / real n = 2 * pi * of_int k" by (auto simp: \_def complex_eq_iff cos_one_2pi_int) with assms have "real n * of_int k = 1" by (simp add: field_simps) also have "real n * of_int k = of_int (int n * k)" by simp also have "1 = (of_int 1 :: real)" by simp also note of_int_eq_iff finally show False using assms by (auto simp: zmult_eq_1_iff) qed have "(\z | z ^ n = 1. z :: complex) = (\k = (\k ^ k)" by (intro sum.cong refl) (auto simp: \_def DeMoivre mult_ac) also have "\ = (\ ^ n - 1) / (\ - 1)" by (subst geometric_sum) auto also have "\ ^ n - 1 = cis (2 * pi) - 1" using assms by (auto simp: \_def DeMoivre) also have "\ = 0" by (simp add: complex_eq_iff) finally show ?thesis by simp qed lemma sum_nth_roots: assumes "n > 1" shows "\{z::complex. z ^ n = c} = 0" proof (cases "c = 0") case True with assms have "{z::complex. z ^ n = c} = {0}" by auto also have "\\ = 0" by simp finally show ?thesis . next case False define c' where "c' = root n (norm c) * cis (arg c / n)" from False and assms have "(\{z. z ^ n = c}) = (\z | z ^ n = 1. c' * z)" by (subst sum.reindex_bij_betw [OF bij_betw_nth_root_unity, symmetric]) (auto simp: sum_distrib_left finite_roots_unity c'_def) also from assms have "\ = 0" by (simp add: sum_distrib_left [symmetric] sum_roots_unity) finally show ?thesis . qed subsection \Square root of complex numbers\ primcorec csqrt :: "complex \ complex" where "Re (csqrt z) = sqrt ((cmod z + Re z) / 2)" | "Im (csqrt z) = (if Im z = 0 then 1 else sgn (Im z)) * sqrt ((cmod z - Re z) / 2)" lemma csqrt_of_real_nonneg [simp]: "Im x = 0 \ Re x \ 0 \ csqrt x = sqrt (Re x)" by (simp add: complex_eq_iff norm_complex_def) lemma csqrt_of_real_nonpos [simp]: "Im x = 0 \ Re x \ 0 \ csqrt x = \ * sqrt \Re x\" by (simp add: complex_eq_iff norm_complex_def) lemma of_real_sqrt: "x \ 0 \ of_real (sqrt x) = csqrt (of_real x)" by (simp add: complex_eq_iff norm_complex_def) lemma csqrt_0 [simp]: "csqrt 0 = 0" by simp lemma csqrt_1 [simp]: "csqrt 1 = 1" by simp lemma csqrt_ii [simp]: "csqrt \ = (1 + \) / sqrt 2" by (simp add: complex_eq_iff Re_divide Im_divide real_sqrt_divide real_div_sqrt) lemma power2_csqrt[simp,algebra]: "(csqrt z)\<^sup>2 = z" proof (cases "Im z = 0") case True then show ?thesis using real_sqrt_pow2[of "Re z"] real_sqrt_pow2[of "- Re z"] by (cases "0::real" "Re z" rule: linorder_cases) (simp_all add: complex_eq_iff Re_power2 Im_power2 power2_eq_square cmod_eq_Re) next case False moreover have "cmod z * cmod z - Re z * Re z = Im z * Im z" by (simp add: norm_complex_def power2_eq_square) moreover have "\Re z\ \ cmod z" by (simp add: norm_complex_def) ultimately show ?thesis by (simp add: Re_power2 Im_power2 complex_eq_iff real_sgn_eq field_simps real_sqrt_mult[symmetric] real_sqrt_divide) qed lemma csqrt_eq_0 [simp]: "csqrt z = 0 \ z = 0" by auto (metis power2_csqrt power_eq_0_iff) lemma csqrt_eq_1 [simp]: "csqrt z = 1 \ z = 1" by auto (metis power2_csqrt power2_eq_1_iff) lemma csqrt_principal: "0 < Re (csqrt z) \ Re (csqrt z) = 0 \ 0 \ Im (csqrt z)" by (auto simp add: not_less cmod_plus_Re_le_0_iff Im_eq_0) lemma Re_csqrt: "0 \ Re (csqrt z)" by (metis csqrt_principal le_less) lemma csqrt_square: assumes "0 < Re b \ (Re b = 0 \ 0 \ Im b)" shows "csqrt (b^2) = b" proof - have "csqrt (b^2) = b \ csqrt (b^2) = - b" by (simp add: power2_eq_iff[symmetric]) moreover have "csqrt (b^2) \ -b \ b = 0" using csqrt_principal[of "b ^ 2"] assms by (intro disjCI notI) (auto simp: complex_eq_iff) ultimately show ?thesis by auto qed lemma csqrt_unique: "w\<^sup>2 = z \ 0 < Re w \ Re w = 0 \ 0 \ Im w \ csqrt z = w" by (auto simp: csqrt_square) lemma csqrt_minus [simp]: assumes "Im x < 0 \ (Im x = 0 \ 0 \ Re x)" shows "csqrt (- x) = \ * csqrt x" proof - have "csqrt ((\ * csqrt x)^2) = \ * csqrt x" proof (rule csqrt_square) have "Im (csqrt x) \ 0" using assms by (auto simp add: cmod_eq_Re mult_le_0_iff field_simps complex_Re_le_cmod) then show "0 < Re (\ * csqrt x) \ Re (\ * csqrt x) = 0 \ 0 \ Im (\ * csqrt x)" by (auto simp add: Re_csqrt simp del: csqrt.simps) qed also have "(\ * csqrt x)^2 = - x" by (simp add: power_mult_distrib) finally show ?thesis . qed text \Legacy theorem names\ lemmas cmod_def = norm_complex_def lemma legacy_Complex_simps: shows Complex_eq_0: "Complex a b = 0 \ a = 0 \ b = 0" and complex_add: "Complex a b + Complex c d = Complex (a + c) (b + d)" and complex_minus: "- (Complex a b) = Complex (- a) (- b)" and complex_diff: "Complex a b - Complex c d = Complex (a - c) (b - d)" and Complex_eq_1: "Complex a b = 1 \ a = 1 \ b = 0" and Complex_eq_neg_1: "Complex a b = - 1 \ a = - 1 \ b = 0" and complex_mult: "Complex a b * Complex c d = Complex (a * c - b * d) (a * d + b * c)" and complex_inverse: "inverse (Complex a b) = Complex (a / (a\<^sup>2 + b\<^sup>2)) (- b / (a\<^sup>2 + b\<^sup>2))" and Complex_eq_numeral: "Complex a b = numeral w \ a = numeral w \ b = 0" and Complex_eq_neg_numeral: "Complex a b = - numeral w \ a = - numeral w \ b = 0" and complex_scaleR: "scaleR r (Complex a b) = Complex (r * a) (r * b)" and Complex_eq_i: "Complex x y = \ \ x = 0 \ y = 1" and i_mult_Complex: "\ * Complex a b = Complex (- b) a" and Complex_mult_i: "Complex a b * \ = Complex (- b) a" and i_complex_of_real: "\ * complex_of_real r = Complex 0 r" and complex_of_real_i: "complex_of_real r * \ = Complex 0 r" and Complex_add_complex_of_real: "Complex x y + complex_of_real r = Complex (x+r) y" and complex_of_real_add_Complex: "complex_of_real r + Complex x y = Complex (r+x) y" and Complex_mult_complex_of_real: "Complex x y * complex_of_real r = Complex (x*r) (y*r)" and complex_of_real_mult_Complex: "complex_of_real r * Complex x y = Complex (r*x) (r*y)" and complex_eq_cancel_iff2: "(Complex x y = complex_of_real xa) = (x = xa \ y = 0)" and complex_cnj: "cnj (Complex a b) = Complex a (- b)" and Complex_sum': "sum (\x. Complex (f x) 0) s = Complex (sum f s) 0" and Complex_sum: "Complex (sum f s) 0 = sum (\x. Complex (f x) 0) s" and complex_of_real_def: "complex_of_real r = Complex r 0" and complex_norm: "cmod (Complex x y) = sqrt (x\<^sup>2 + y\<^sup>2)" by (simp_all add: norm_complex_def field_simps complex_eq_iff Re_divide Im_divide) lemma Complex_in_Reals: "Complex x 0 \ \" by (metis Reals_of_real complex_of_real_def) end diff --git a/src/HOL/Computational_Algebra/Polynomial.thy b/src/HOL/Computational_Algebra/Polynomial.thy --- a/src/HOL/Computational_Algebra/Polynomial.thy +++ b/src/HOL/Computational_Algebra/Polynomial.thy @@ -1,4708 +1,5019 @@ (* Title: HOL/Computational_Algebra/Polynomial.thy Author: Brian Huffman Author: Clemens Ballarin Author: Amine Chaieb Author: Florian Haftmann *) section \Polynomials as type over a ring structure\ theory Polynomial imports Complex_Main "HOL-Library.More_List" "HOL-Library.Infinite_Set" Factorial_Ring begin subsection \Auxiliary: operations for lists (later) representing coefficients\ definition cCons :: "'a::zero \ 'a list \ 'a list" (infixr "##" 65) where "x ## xs = (if xs = [] \ x = 0 then [] else x # xs)" lemma cCons_0_Nil_eq [simp]: "0 ## [] = []" by (simp add: cCons_def) lemma cCons_Cons_eq [simp]: "x ## y # ys = x # y # ys" by (simp add: cCons_def) lemma cCons_append_Cons_eq [simp]: "x ## xs @ y # ys = x # xs @ y # ys" by (simp add: cCons_def) lemma cCons_not_0_eq [simp]: "x \ 0 \ x ## xs = x # xs" by (simp add: cCons_def) lemma strip_while_not_0_Cons_eq [simp]: "strip_while (\x. x = 0) (x # xs) = x ## strip_while (\x. x = 0) xs" proof (cases "x = 0") case False then show ?thesis by simp next case True show ?thesis proof (induct xs rule: rev_induct) case Nil with True show ?case by simp next case (snoc y ys) then show ?case by (cases "y = 0") (simp_all add: append_Cons [symmetric] del: append_Cons) qed qed lemma tl_cCons [simp]: "tl (x ## xs) = xs" by (simp add: cCons_def) subsection \Definition of type \poly\\ typedef (overloaded) 'a poly = "{f :: nat \ 'a::zero. \\<^sub>\ n. f n = 0}" morphisms coeff Abs_poly by (auto intro!: ALL_MOST) setup_lifting type_definition_poly lemma poly_eq_iff: "p = q \ (\n. coeff p n = coeff q n)" by (simp add: coeff_inject [symmetric] fun_eq_iff) lemma poly_eqI: "(\n. coeff p n = coeff q n) \ p = q" by (simp add: poly_eq_iff) lemma MOST_coeff_eq_0: "\\<^sub>\ n. coeff p n = 0" using coeff [of p] by simp subsection \Degree of a polynomial\ definition degree :: "'a::zero poly \ nat" where "degree p = (LEAST n. \i>n. coeff p i = 0)" lemma coeff_eq_0: assumes "degree p < n" shows "coeff p n = 0" proof - have "\n. \i>n. coeff p i = 0" using MOST_coeff_eq_0 by (simp add: MOST_nat) then have "\i>degree p. coeff p i = 0" unfolding degree_def by (rule LeastI_ex) with assms show ?thesis by simp qed lemma le_degree: "coeff p n \ 0 \ n \ degree p" by (erule contrapos_np, rule coeff_eq_0, simp) lemma degree_le: "\i>n. coeff p i = 0 \ degree p \ n" unfolding degree_def by (erule Least_le) lemma less_degree_imp: "n < degree p \ \i>n. coeff p i \ 0" unfolding degree_def by (drule not_less_Least, simp) subsection \The zero polynomial\ instantiation poly :: (zero) zero begin lift_definition zero_poly :: "'a poly" is "\_. 0" by (rule MOST_I) simp instance .. end lemma coeff_0 [simp]: "coeff 0 n = 0" by transfer rule lemma degree_0 [simp]: "degree 0 = 0" by (rule order_antisym [OF degree_le le0]) simp lemma leading_coeff_neq_0: assumes "p \ 0" shows "coeff p (degree p) \ 0" proof (cases "degree p") case 0 from \p \ 0\ obtain n where "coeff p n \ 0" by (auto simp add: poly_eq_iff) then have "n \ degree p" by (rule le_degree) with \coeff p n \ 0\ and \degree p = 0\ show "coeff p (degree p) \ 0" by simp next case (Suc n) from \degree p = Suc n\ have "n < degree p" by simp then have "\i>n. coeff p i \ 0" by (rule less_degree_imp) then obtain i where "n < i" and "coeff p i \ 0" by blast from \degree p = Suc n\ and \n < i\ have "degree p \ i" by simp also from \coeff p i \ 0\ have "i \ degree p" by (rule le_degree) finally have "degree p = i" . with \coeff p i \ 0\ show "coeff p (degree p) \ 0" by simp qed lemma leading_coeff_0_iff [simp]: "coeff p (degree p) = 0 \ p = 0" by (cases "p = 0") (simp_all add: leading_coeff_neq_0) lemma eq_zero_or_degree_less: assumes "degree p \ n" and "coeff p n = 0" shows "p = 0 \ degree p < n" proof (cases n) case 0 with \degree p \ n\ and \coeff p n = 0\ have "coeff p (degree p) = 0" by simp then have "p = 0" by simp then show ?thesis .. next case (Suc m) from \degree p \ n\ have "\i>n. coeff p i = 0" by (simp add: coeff_eq_0) with \coeff p n = 0\ have "\i\n. coeff p i = 0" by (simp add: le_less) with \n = Suc m\ have "\i>m. coeff p i = 0" by (simp add: less_eq_Suc_le) then have "degree p \ m" by (rule degree_le) with \n = Suc m\ have "degree p < n" by (simp add: less_Suc_eq_le) then show ?thesis .. qed lemma coeff_0_degree_minus_1: "coeff rrr dr = 0 \ degree rrr \ dr \ degree rrr \ dr - 1" using eq_zero_or_degree_less by fastforce subsection \List-style constructor for polynomials\ lift_definition pCons :: "'a::zero \ 'a poly \ 'a poly" is "\a p. case_nat a (coeff p)" by (rule MOST_SucD) (simp add: MOST_coeff_eq_0) lemmas coeff_pCons = pCons.rep_eq lemma coeff_pCons_0 [simp]: "coeff (pCons a p) 0 = a" by transfer simp lemma coeff_pCons_Suc [simp]: "coeff (pCons a p) (Suc n) = coeff p n" by (simp add: coeff_pCons) lemma degree_pCons_le: "degree (pCons a p) \ Suc (degree p)" by (rule degree_le) (simp add: coeff_eq_0 coeff_pCons split: nat.split) lemma degree_pCons_eq: "p \ 0 \ degree (pCons a p) = Suc (degree p)" by (simp add: degree_pCons_le le_antisym le_degree) lemma degree_pCons_0: "degree (pCons a 0) = 0" proof - have "degree (pCons a 0) \ Suc 0" by (metis (no_types) degree_0 degree_pCons_le) then show ?thesis by (metis coeff_0 coeff_pCons_Suc degree_0 eq_zero_or_degree_less less_Suc0) qed lemma degree_pCons_eq_if [simp]: "degree (pCons a p) = (if p = 0 then 0 else Suc (degree p))" by (simp add: degree_pCons_0 degree_pCons_eq) lemma pCons_0_0 [simp]: "pCons 0 0 = 0" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma pCons_eq_iff [simp]: "pCons a p = pCons b q \ a = b \ p = q" proof safe assume "pCons a p = pCons b q" then have "coeff (pCons a p) 0 = coeff (pCons b q) 0" by simp then show "a = b" by simp next assume "pCons a p = pCons b q" then have "coeff (pCons a p) (Suc n) = coeff (pCons b q) (Suc n)" for n by simp then show "p = q" by (simp add: poly_eq_iff) qed lemma pCons_eq_0_iff [simp]: "pCons a p = 0 \ a = 0 \ p = 0" using pCons_eq_iff [of a p 0 0] by simp lemma pCons_cases [cases type: poly]: obtains (pCons) a q where "p = pCons a q" proof show "p = pCons (coeff p 0) (Abs_poly (\n. coeff p (Suc n)))" by transfer (simp_all add: MOST_inj[where f=Suc and P="\n. p n = 0" for p] fun_eq_iff Abs_poly_inverse split: nat.split) qed lemma pCons_induct [case_names 0 pCons, induct type: poly]: assumes zero: "P 0" assumes pCons: "\a p. a \ 0 \ p \ 0 \ P p \ P (pCons a p)" shows "P p" proof (induct p rule: measure_induct_rule [where f=degree]) case (less p) obtain a q where "p = pCons a q" by (rule pCons_cases) have "P q" proof (cases "q = 0") case True then show "P q" by (simp add: zero) next case False then have "degree (pCons a q) = Suc (degree q)" by (rule degree_pCons_eq) with \p = pCons a q\ have "degree q < degree p" by simp then show "P q" by (rule less.hyps) qed have "P (pCons a q)" proof (cases "a \ 0 \ q \ 0") case True with \P q\ show ?thesis by (auto intro: pCons) next case False with zero show ?thesis by simp qed with \p = pCons a q\ show ?case by simp qed lemma degree_eq_zeroE: fixes p :: "'a::zero poly" assumes "degree p = 0" obtains a where "p = pCons a 0" proof - obtain a q where p: "p = pCons a q" by (cases p) with assms have "q = 0" by (cases "q = 0") simp_all with p have "p = pCons a 0" by simp then show thesis .. qed subsection \Quickcheck generator for polynomials\ quickcheck_generator poly constructors: "0 :: _ poly", pCons subsection \List-style syntax for polynomials\ syntax "_poly" :: "args \ 'a poly" ("[:(_):]") translations "[:x, xs:]" \ "CONST pCons x [:xs:]" "[:x:]" \ "CONST pCons x 0" "[:x:]" \ "CONST pCons x (_constrain 0 t)" subsection \Representation of polynomials by lists of coefficients\ primrec Poly :: "'a::zero list \ 'a poly" where [code_post]: "Poly [] = 0" | [code_post]: "Poly (a # as) = pCons a (Poly as)" lemma Poly_replicate_0 [simp]: "Poly (replicate n 0) = 0" by (induct n) simp_all lemma Poly_eq_0: "Poly as = 0 \ (\n. as = replicate n 0)" by (induct as) (auto simp add: Cons_replicate_eq) lemma Poly_append_replicate_zero [simp]: "Poly (as @ replicate n 0) = Poly as" by (induct as) simp_all lemma Poly_snoc_zero [simp]: "Poly (as @ [0]) = Poly as" using Poly_append_replicate_zero [of as 1] by simp lemma Poly_cCons_eq_pCons_Poly [simp]: "Poly (a ## p) = pCons a (Poly p)" by (simp add: cCons_def) lemma Poly_on_rev_starting_with_0 [simp]: "hd as = 0 \ Poly (rev (tl as)) = Poly (rev as)" by (cases as) simp_all lemma degree_Poly: "degree (Poly xs) \ length xs" by (induct xs) simp_all lemma coeff_Poly_eq [simp]: "coeff (Poly xs) = nth_default 0 xs" by (induct xs) (simp_all add: fun_eq_iff coeff_pCons split: nat.splits) definition coeffs :: "'a poly \ 'a::zero list" where "coeffs p = (if p = 0 then [] else map (\i. coeff p i) [0 ..< Suc (degree p)])" lemma coeffs_eq_Nil [simp]: "coeffs p = [] \ p = 0" by (simp add: coeffs_def) lemma not_0_coeffs_not_Nil: "p \ 0 \ coeffs p \ []" by simp lemma coeffs_0_eq_Nil [simp]: "coeffs 0 = []" by simp lemma coeffs_pCons_eq_cCons [simp]: "coeffs (pCons a p) = a ## coeffs p" proof - have *: "\m\set ms. m > 0 \ map (case_nat x f) ms = map f (map (\n. n - 1) ms)" for ms :: "nat list" and f :: "nat \ 'a" and x :: "'a" by (induct ms) (auto split: nat.split) show ?thesis by (simp add: * coeffs_def upt_conv_Cons coeff_pCons map_decr_upt del: upt_Suc) qed lemma length_coeffs: "p \ 0 \ length (coeffs p) = degree p + 1" by (simp add: coeffs_def) lemma coeffs_nth: "p \ 0 \ n \ degree p \ coeffs p ! n = coeff p n" by (auto simp: coeffs_def simp del: upt_Suc) lemma coeff_in_coeffs: "p \ 0 \ n \ degree p \ coeff p n \ set (coeffs p)" using coeffs_nth [of p n, symmetric] by (simp add: length_coeffs) lemma not_0_cCons_eq [simp]: "p \ 0 \ a ## coeffs p = a # coeffs p" by (simp add: cCons_def) lemma Poly_coeffs [simp, code abstype]: "Poly (coeffs p) = p" by (induct p) auto lemma coeffs_Poly [simp]: "coeffs (Poly as) = strip_while (HOL.eq 0) as" proof (induct as) case Nil then show ?case by simp next case (Cons a as) from replicate_length_same [of as 0] have "(\n. as \ replicate n 0) \ (\a\set as. a \ 0)" by (auto dest: sym [of _ as]) with Cons show ?case by auto qed lemma no_trailing_coeffs [simp]: "no_trailing (HOL.eq 0) (coeffs p)" by (induct p) auto lemma strip_while_coeffs [simp]: "strip_while (HOL.eq 0) (coeffs p) = coeffs p" by simp lemma coeffs_eq_iff: "p = q \ coeffs p = coeffs q" (is "?P \ ?Q") proof assume ?P then show ?Q by simp next assume ?Q then have "Poly (coeffs p) = Poly (coeffs q)" by simp then show ?P by simp qed lemma nth_default_coeffs_eq: "nth_default 0 (coeffs p) = coeff p" by (simp add: fun_eq_iff coeff_Poly_eq [symmetric]) lemma [code]: "coeff p = nth_default 0 (coeffs p)" by (simp add: nth_default_coeffs_eq) lemma coeffs_eqI: assumes coeff: "\n. coeff p n = nth_default 0 xs n" assumes zero: "no_trailing (HOL.eq 0) xs" shows "coeffs p = xs" proof - from coeff have "p = Poly xs" by (simp add: poly_eq_iff) with zero show ?thesis by simp qed lemma degree_eq_length_coeffs [code]: "degree p = length (coeffs p) - 1" by (simp add: coeffs_def) lemma length_coeffs_degree: "p \ 0 \ length (coeffs p) = Suc (degree p)" by (induct p) (auto simp: cCons_def) lemma [code abstract]: "coeffs 0 = []" by (fact coeffs_0_eq_Nil) lemma [code abstract]: "coeffs (pCons a p) = a ## coeffs p" by (fact coeffs_pCons_eq_cCons) lemma set_coeffs_subset_singleton_0_iff [simp]: "set (coeffs p) \ {0} \ p = 0" by (auto simp add: coeffs_def intro: classical) lemma set_coeffs_not_only_0 [simp]: "set (coeffs p) \ {0}" by (auto simp add: set_eq_subset) lemma forall_coeffs_conv: "(\n. P (coeff p n)) \ (\c \ set (coeffs p). P c)" if "P 0" using that by (auto simp add: coeffs_def) (metis atLeastLessThan_iff coeff_eq_0 not_less_iff_gr_or_eq zero_le) instantiation poly :: ("{zero, equal}") equal begin definition [code]: "HOL.equal (p::'a poly) q \ HOL.equal (coeffs p) (coeffs q)" instance by standard (simp add: equal equal_poly_def coeffs_eq_iff) end lemma [code nbe]: "HOL.equal (p :: _ poly) p \ True" by (fact equal_refl) definition is_zero :: "'a::zero poly \ bool" where [code]: "is_zero p \ List.null (coeffs p)" lemma is_zero_null [code_abbrev]: "is_zero p \ p = 0" by (simp add: is_zero_def null_def) subsubsection \Reconstructing the polynomial from the list\ \ \contributed by Sebastiaan J.C. Joosten and René Thiemann\ definition poly_of_list :: "'a::comm_monoid_add list \ 'a poly" where [simp]: "poly_of_list = Poly" lemma poly_of_list_impl [code abstract]: "coeffs (poly_of_list as) = strip_while (HOL.eq 0) as" by simp subsection \Fold combinator for polynomials\ definition fold_coeffs :: "('a::zero \ 'b \ 'b) \ 'a poly \ 'b \ 'b" where "fold_coeffs f p = foldr f (coeffs p)" lemma fold_coeffs_0_eq [simp]: "fold_coeffs f 0 = id" by (simp add: fold_coeffs_def) lemma fold_coeffs_pCons_eq [simp]: "f 0 = id \ fold_coeffs f (pCons a p) = f a \ fold_coeffs f p" by (simp add: fold_coeffs_def cCons_def fun_eq_iff) lemma fold_coeffs_pCons_0_0_eq [simp]: "fold_coeffs f (pCons 0 0) = id" by (simp add: fold_coeffs_def) lemma fold_coeffs_pCons_coeff_not_0_eq [simp]: "a \ 0 \ fold_coeffs f (pCons a p) = f a \ fold_coeffs f p" by (simp add: fold_coeffs_def) lemma fold_coeffs_pCons_not_0_0_eq [simp]: "p \ 0 \ fold_coeffs f (pCons a p) = f a \ fold_coeffs f p" by (simp add: fold_coeffs_def) subsection \Canonical morphism on polynomials -- evaluation\ definition poly :: \'a::comm_semiring_0 poly \ 'a \ 'a\ where \poly p a = horner_sum id a (coeffs p)\ lemma poly_eq_fold_coeffs: \poly p = fold_coeffs (\a f x. a + x * f x) p (\x. 0)\ by (induction p) (auto simp add: fun_eq_iff poly_def) lemma poly_0 [simp]: "poly 0 x = 0" by (simp add: poly_def) lemma poly_pCons [simp]: "poly (pCons a p) x = a + x * poly p x" by (cases "p = 0 \ a = 0") (auto simp add: poly_def) lemma poly_altdef: "poly p x = (\i\degree p. coeff p i * x ^ i)" for x :: "'a::{comm_semiring_0,semiring_1}" proof (induction p rule: pCons_induct) case 0 then show ?case by simp next case (pCons a p) show ?case proof (cases "p = 0") case True then show ?thesis by simp next case False let ?p' = "pCons a p" note poly_pCons[of a p x] also note pCons.IH also have "a + x * (\i\degree p. coeff p i * x ^ i) = coeff ?p' 0 * x^0 + (\i\degree p. coeff ?p' (Suc i) * x^Suc i)" by (simp add: field_simps sum_distrib_left coeff_pCons) also note sum.atMost_Suc_shift[symmetric] also note degree_pCons_eq[OF \p \ 0\, of a, symmetric] finally show ?thesis . qed qed lemma poly_0_coeff_0: "poly p 0 = coeff p 0" by (cases p) (auto simp: poly_altdef) subsection \Monomials\ lift_definition monom :: "'a \ nat \ 'a::zero poly" is "\a m n. if m = n then a else 0" by (simp add: MOST_iff_cofinite) lemma coeff_monom [simp]: "coeff (monom a m) n = (if m = n then a else 0)" by transfer rule lemma monom_0: "monom a 0 = pCons a 0" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma monom_Suc: "monom a (Suc n) = pCons 0 (monom a n)" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma monom_eq_0 [simp]: "monom 0 n = 0" by (rule poly_eqI) simp lemma monom_eq_0_iff [simp]: "monom a n = 0 \ a = 0" by (simp add: poly_eq_iff) lemma monom_eq_iff [simp]: "monom a n = monom b n \ a = b" by (simp add: poly_eq_iff) lemma degree_monom_le: "degree (monom a n) \ n" by (rule degree_le, simp) lemma degree_monom_eq: "a \ 0 \ degree (monom a n) = n" by (metis coeff_monom leading_coeff_0_iff) lemma coeffs_monom [code abstract]: "coeffs (monom a n) = (if a = 0 then [] else replicate n 0 @ [a])" by (induct n) (simp_all add: monom_0 monom_Suc) lemma fold_coeffs_monom [simp]: "a \ 0 \ fold_coeffs f (monom a n) = f 0 ^^ n \ f a" by (simp add: fold_coeffs_def coeffs_monom fun_eq_iff) lemma poly_monom: "poly (monom a n) x = a * x ^ n" for a x :: "'a::comm_semiring_1" by (cases "a = 0", simp_all) (induct n, simp_all add: mult.left_commute poly_eq_fold_coeffs) lemma monom_eq_iff': "monom c n = monom d m \ c = d \ (c = 0 \ n = m)" by (auto simp: poly_eq_iff) lemma monom_eq_const_iff: "monom c n = [:d:] \ c = d \ (c = 0 \ n = 0)" using monom_eq_iff'[of c n d 0] by (simp add: monom_0) subsection \Leading coefficient\ abbreviation lead_coeff:: "'a::zero poly \ 'a" where "lead_coeff p \ coeff p (degree p)" lemma lead_coeff_pCons[simp]: "p \ 0 \ lead_coeff (pCons a p) = lead_coeff p" "p = 0 \ lead_coeff (pCons a p) = a" by auto lemma lead_coeff_monom [simp]: "lead_coeff (monom c n) = c" by (cases "c = 0") (simp_all add: degree_monom_eq) lemma last_coeffs_eq_coeff_degree: "last (coeffs p) = lead_coeff p" if "p \ 0" using that by (simp add: coeffs_def) subsection \Addition and subtraction\ instantiation poly :: (comm_monoid_add) comm_monoid_add begin lift_definition plus_poly :: "'a poly \ 'a poly \ 'a poly" is "\p q n. coeff p n + coeff q n" proof - fix q p :: "'a poly" show "\\<^sub>\n. coeff p n + coeff q n = 0" using MOST_coeff_eq_0[of p] MOST_coeff_eq_0[of q] by eventually_elim simp qed lemma coeff_add [simp]: "coeff (p + q) n = coeff p n + coeff q n" by (simp add: plus_poly.rep_eq) instance proof fix p q r :: "'a poly" show "(p + q) + r = p + (q + r)" by (simp add: poly_eq_iff add.assoc) show "p + q = q + p" by (simp add: poly_eq_iff add.commute) show "0 + p = p" by (simp add: poly_eq_iff) qed end instantiation poly :: (cancel_comm_monoid_add) cancel_comm_monoid_add begin lift_definition minus_poly :: "'a poly \ 'a poly \ 'a poly" is "\p q n. coeff p n - coeff q n" proof - fix q p :: "'a poly" show "\\<^sub>\n. coeff p n - coeff q n = 0" using MOST_coeff_eq_0[of p] MOST_coeff_eq_0[of q] by eventually_elim simp qed lemma coeff_diff [simp]: "coeff (p - q) n = coeff p n - coeff q n" by (simp add: minus_poly.rep_eq) instance proof fix p q r :: "'a poly" show "p + q - p = q" by (simp add: poly_eq_iff) show "p - q - r = p - (q + r)" by (simp add: poly_eq_iff diff_diff_eq) qed end instantiation poly :: (ab_group_add) ab_group_add begin lift_definition uminus_poly :: "'a poly \ 'a poly" is "\p n. - coeff p n" proof - fix p :: "'a poly" show "\\<^sub>\n. - coeff p n = 0" using MOST_coeff_eq_0 by simp qed lemma coeff_minus [simp]: "coeff (- p) n = - coeff p n" by (simp add: uminus_poly.rep_eq) instance proof fix p q :: "'a poly" show "- p + p = 0" by (simp add: poly_eq_iff) show "p - q = p + - q" by (simp add: poly_eq_iff) qed end lemma add_pCons [simp]: "pCons a p + pCons b q = pCons (a + b) (p + q)" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma minus_pCons [simp]: "- pCons a p = pCons (- a) (- p)" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma diff_pCons [simp]: "pCons a p - pCons b q = pCons (a - b) (p - q)" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma degree_add_le_max: "degree (p + q) \ max (degree p) (degree q)" by (rule degree_le) (auto simp add: coeff_eq_0) lemma degree_add_le: "degree p \ n \ degree q \ n \ degree (p + q) \ n" by (auto intro: order_trans degree_add_le_max) lemma degree_add_less: "degree p < n \ degree q < n \ degree (p + q) < n" by (auto intro: le_less_trans degree_add_le_max) lemma degree_add_eq_right: assumes "degree p < degree q" shows "degree (p + q) = degree q" proof (cases "q = 0") case False show ?thesis proof (rule order_antisym) show "degree (p + q) \ degree q" by (simp add: assms degree_add_le order.strict_implies_order) show "degree q \ degree (p + q)" by (simp add: False assms coeff_eq_0 le_degree) qed qed (use assms in auto) lemma degree_add_eq_left: "degree q < degree p \ degree (p + q) = degree p" using degree_add_eq_right [of q p] by (simp add: add.commute) lemma degree_minus [simp]: "degree (- p) = degree p" by (simp add: degree_def) lemma lead_coeff_add_le: "degree p < degree q \ lead_coeff (p + q) = lead_coeff q" by (metis coeff_add coeff_eq_0 monoid_add_class.add.left_neutral degree_add_eq_right) lemma lead_coeff_minus: "lead_coeff (- p) = - lead_coeff p" by (metis coeff_minus degree_minus) lemma degree_diff_le_max: "degree (p - q) \ max (degree p) (degree q)" for p q :: "'a::ab_group_add poly" using degree_add_le [where p=p and q="-q"] by simp lemma degree_diff_le: "degree p \ n \ degree q \ n \ degree (p - q) \ n" for p q :: "'a::ab_group_add poly" using degree_add_le [of p n "- q"] by simp lemma degree_diff_less: "degree p < n \ degree q < n \ degree (p - q) < n" for p q :: "'a::ab_group_add poly" using degree_add_less [of p n "- q"] by simp lemma add_monom: "monom a n + monom b n = monom (a + b) n" by (rule poly_eqI) simp lemma diff_monom: "monom a n - monom b n = monom (a - b) n" by (rule poly_eqI) simp lemma minus_monom: "- monom a n = monom (- a) n" by (rule poly_eqI) simp lemma coeff_sum: "coeff (\x\A. p x) i = (\x\A. coeff (p x) i)" by (induct A rule: infinite_finite_induct) simp_all lemma monom_sum: "monom (\x\A. a x) n = (\x\A. monom (a x) n)" by (rule poly_eqI) (simp add: coeff_sum) fun plus_coeffs :: "'a::comm_monoid_add list \ 'a list \ 'a list" where "plus_coeffs xs [] = xs" | "plus_coeffs [] ys = ys" | "plus_coeffs (x # xs) (y # ys) = (x + y) ## plus_coeffs xs ys" lemma coeffs_plus_eq_plus_coeffs [code abstract]: "coeffs (p + q) = plus_coeffs (coeffs p) (coeffs q)" proof - have *: "nth_default 0 (plus_coeffs xs ys) n = nth_default 0 xs n + nth_default 0 ys n" for xs ys :: "'a list" and n proof (induct xs ys arbitrary: n rule: plus_coeffs.induct) case (3 x xs y ys n) then show ?case by (cases n) (auto simp add: cCons_def) qed simp_all have **: "no_trailing (HOL.eq 0) (plus_coeffs xs ys)" if "no_trailing (HOL.eq 0) xs" and "no_trailing (HOL.eq 0) ys" for xs ys :: "'a list" using that by (induct xs ys rule: plus_coeffs.induct) (simp_all add: cCons_def) show ?thesis by (rule coeffs_eqI) (auto simp add: * nth_default_coeffs_eq intro: **) qed lemma coeffs_uminus [code abstract]: "coeffs (- p) = map uminus (coeffs p)" proof - have eq_0: "HOL.eq 0 \ uminus = HOL.eq (0::'a)" by (simp add: fun_eq_iff) show ?thesis by (rule coeffs_eqI) (simp_all add: nth_default_map_eq nth_default_coeffs_eq no_trailing_map eq_0) qed lemma [code]: "p - q = p + - q" for p q :: "'a::ab_group_add poly" by (fact diff_conv_add_uminus) lemma poly_add [simp]: "poly (p + q) x = poly p x + poly q x" proof (induction p arbitrary: q) case (pCons a p) then show ?case by (cases q) (simp add: algebra_simps) qed auto lemma poly_minus [simp]: "poly (- p) x = - poly p x" for x :: "'a::comm_ring" by (induct p) simp_all lemma poly_diff [simp]: "poly (p - q) x = poly p x - poly q x" for x :: "'a::comm_ring" using poly_add [of p "- q" x] by simp lemma poly_sum: "poly (\k\A. p k) x = (\k\A. poly (p k) x)" by (induct A rule: infinite_finite_induct) simp_all lemma degree_sum_le: "finite S \ (\p. p \ S \ degree (f p) \ n) \ degree (sum f S) \ n" proof (induct S rule: finite_induct) case empty then show ?case by simp next case (insert p S) then have "degree (sum f S) \ n" "degree (f p) \ n" by auto then show ?case unfolding sum.insert[OF insert(1-2)] by (metis degree_add_le) qed lemma poly_as_sum_of_monoms': assumes "degree p \ n" shows "(\i\n. monom (coeff p i) i) = p" proof - have eq: "\i. {..n} \ {i} = (if i \ n then {i} else {})" by auto from assms show ?thesis by (simp add: poly_eq_iff coeff_sum coeff_eq_0 sum.If_cases eq if_distrib[where f="\x. x * a" for a]) qed lemma poly_as_sum_of_monoms: "(\i\degree p. monom (coeff p i) i) = p" by (intro poly_as_sum_of_monoms' order_refl) lemma Poly_snoc: "Poly (xs @ [x]) = Poly xs + monom x (length xs)" by (induct xs) (simp_all add: monom_0 monom_Suc) subsection \Multiplication by a constant, polynomial multiplication and the unit polynomial\ lift_definition smult :: "'a::comm_semiring_0 \ 'a poly \ 'a poly" is "\a p n. a * coeff p n" proof - fix a :: 'a and p :: "'a poly" show "\\<^sub>\ i. a * coeff p i = 0" using MOST_coeff_eq_0[of p] by eventually_elim simp qed lemma coeff_smult [simp]: "coeff (smult a p) n = a * coeff p n" by (simp add: smult.rep_eq) lemma degree_smult_le: "degree (smult a p) \ degree p" by (rule degree_le) (simp add: coeff_eq_0) lemma smult_smult [simp]: "smult a (smult b p) = smult (a * b) p" by (rule poly_eqI) (simp add: mult.assoc) lemma smult_0_right [simp]: "smult a 0 = 0" by (rule poly_eqI) simp lemma smult_0_left [simp]: "smult 0 p = 0" by (rule poly_eqI) simp lemma smult_1_left [simp]: "smult (1::'a::comm_semiring_1) p = p" by (rule poly_eqI) simp lemma smult_add_right: "smult a (p + q) = smult a p + smult a q" by (rule poly_eqI) (simp add: algebra_simps) lemma smult_add_left: "smult (a + b) p = smult a p + smult b p" by (rule poly_eqI) (simp add: algebra_simps) lemma smult_minus_right [simp]: "smult a (- p) = - smult a p" for a :: "'a::comm_ring" by (rule poly_eqI) simp lemma smult_minus_left [simp]: "smult (- a) p = - smult a p" for a :: "'a::comm_ring" by (rule poly_eqI) simp lemma smult_diff_right: "smult a (p - q) = smult a p - smult a q" for a :: "'a::comm_ring" by (rule poly_eqI) (simp add: algebra_simps) lemma smult_diff_left: "smult (a - b) p = smult a p - smult b p" for a b :: "'a::comm_ring" by (rule poly_eqI) (simp add: algebra_simps) lemmas smult_distribs = smult_add_left smult_add_right smult_diff_left smult_diff_right lemma smult_pCons [simp]: "smult a (pCons b p) = pCons (a * b) (smult a p)" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma smult_monom: "smult a (monom b n) = monom (a * b) n" by (induct n) (simp_all add: monom_0 monom_Suc) lemma smult_Poly: "smult c (Poly xs) = Poly (map ((*) c) xs)" by (auto simp: poly_eq_iff nth_default_def) lemma degree_smult_eq [simp]: "degree (smult a p) = (if a = 0 then 0 else degree p)" for a :: "'a::{comm_semiring_0,semiring_no_zero_divisors}" by (cases "a = 0") (simp_all add: degree_def) lemma smult_eq_0_iff [simp]: "smult a p = 0 \ a = 0 \ p = 0" for a :: "'a::{comm_semiring_0,semiring_no_zero_divisors}" by (simp add: poly_eq_iff) lemma coeffs_smult [code abstract]: "coeffs (smult a p) = (if a = 0 then [] else map (Groups.times a) (coeffs p))" for p :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" proof - have eq_0: "HOL.eq 0 \ times a = HOL.eq (0::'a)" if "a \ 0" using that by (simp add: fun_eq_iff) show ?thesis by (rule coeffs_eqI) (auto simp add: no_trailing_map nth_default_map_eq nth_default_coeffs_eq eq_0) qed lemma smult_eq_iff: fixes b :: "'a :: field" assumes "b \ 0" shows "smult a p = smult b q \ smult (a / b) p = q" (is "?lhs \ ?rhs") proof assume ?lhs also from assms have "smult (inverse b) \ = q" by simp finally show ?rhs by (simp add: field_simps) next assume ?rhs with assms show ?lhs by auto qed instantiation poly :: (comm_semiring_0) comm_semiring_0 begin definition "p * q = fold_coeffs (\a p. smult a q + pCons 0 p) p 0" lemma mult_poly_0_left: "(0::'a poly) * q = 0" by (simp add: times_poly_def) lemma mult_pCons_left [simp]: "pCons a p * q = smult a q + pCons 0 (p * q)" by (cases "p = 0 \ a = 0") (auto simp add: times_poly_def) lemma mult_poly_0_right: "p * (0::'a poly) = 0" by (induct p) (simp_all add: mult_poly_0_left) lemma mult_pCons_right [simp]: "p * pCons a q = smult a p + pCons 0 (p * q)" by (induct p) (simp_all add: mult_poly_0_left algebra_simps) lemmas mult_poly_0 = mult_poly_0_left mult_poly_0_right lemma mult_smult_left [simp]: "smult a p * q = smult a (p * q)" by (induct p) (simp_all add: mult_poly_0 smult_add_right) lemma mult_smult_right [simp]: "p * smult a q = smult a (p * q)" by (induct q) (simp_all add: mult_poly_0 smult_add_right) lemma mult_poly_add_left: "(p + q) * r = p * r + q * r" for p q r :: "'a poly" by (induct r) (simp_all add: mult_poly_0 smult_distribs algebra_simps) instance proof fix p q r :: "'a poly" show 0: "0 * p = 0" by (rule mult_poly_0_left) show "p * 0 = 0" by (rule mult_poly_0_right) show "(p + q) * r = p * r + q * r" by (rule mult_poly_add_left) show "(p * q) * r = p * (q * r)" by (induct p) (simp_all add: mult_poly_0 mult_poly_add_left) show "p * q = q * p" by (induct p) (simp_all add: mult_poly_0) qed end lemma coeff_mult_degree_sum: "coeff (p * q) (degree p + degree q) = coeff p (degree p) * coeff q (degree q)" by (induct p) (simp_all add: coeff_eq_0) instance poly :: ("{comm_semiring_0,semiring_no_zero_divisors}") semiring_no_zero_divisors proof fix p q :: "'a poly" assume "p \ 0" and "q \ 0" have "coeff (p * q) (degree p + degree q) = coeff p (degree p) * coeff q (degree q)" by (rule coeff_mult_degree_sum) also from \p \ 0\ \q \ 0\ have "coeff p (degree p) * coeff q (degree q) \ 0" by simp finally have "\n. coeff (p * q) n \ 0" .. then show "p * q \ 0" by (simp add: poly_eq_iff) qed instance poly :: (comm_semiring_0_cancel) comm_semiring_0_cancel .. lemma coeff_mult: "coeff (p * q) n = (\i\n. coeff p i * coeff q (n-i))" proof (induct p arbitrary: n) case 0 show ?case by simp next case (pCons a p n) then show ?case by (cases n) (simp_all add: sum.atMost_Suc_shift del: sum.atMost_Suc) qed lemma degree_mult_le: "degree (p * q) \ degree p + degree q" proof (rule degree_le) show "\i>degree p + degree q. coeff (p * q) i = 0" by (induct p) (simp_all add: coeff_eq_0 coeff_pCons split: nat.split) qed lemma mult_monom: "monom a m * monom b n = monom (a * b) (m + n)" by (induct m) (simp add: monom_0 smult_monom, simp add: monom_Suc) instantiation poly :: (comm_semiring_1) comm_semiring_1 begin lift_definition one_poly :: "'a poly" is "\n. of_bool (n = 0)" by (rule MOST_SucD) simp lemma coeff_1 [simp]: "coeff 1 n = of_bool (n = 0)" by (simp add: one_poly.rep_eq) lemma one_pCons: "1 = [:1:]" by (simp add: poly_eq_iff coeff_pCons split: nat.splits) lemma pCons_one: "[:1:] = 1" by (simp add: one_pCons) instance by standard (simp_all add: one_pCons) end lemma poly_1 [simp]: "poly 1 x = 1" by (simp add: one_pCons) lemma one_poly_eq_simps [simp]: "1 = [:1:] \ True" "[:1:] = 1 \ True" by (simp_all add: one_pCons) lemma degree_1 [simp]: "degree 1 = 0" by (simp add: one_pCons) lemma coeffs_1_eq [simp, code abstract]: "coeffs 1 = [1]" by (simp add: one_pCons) lemma smult_one [simp]: "smult c 1 = [:c:]" by (simp add: one_pCons) lemma monom_eq_1 [simp]: "monom 1 0 = 1" by (simp add: monom_0 one_pCons) lemma monom_eq_1_iff: "monom c n = 1 \ c = 1 \ n = 0" using monom_eq_const_iff [of c n 1] by auto lemma monom_altdef: "monom c n = smult c ([:0, 1:] ^ n)" by (induct n) (simp_all add: monom_0 monom_Suc) instance poly :: ("{comm_semiring_1,semiring_1_no_zero_divisors}") semiring_1_no_zero_divisors .. instance poly :: (comm_ring) comm_ring .. instance poly :: (comm_ring_1) comm_ring_1 .. instance poly :: (comm_ring_1) comm_semiring_1_cancel .. lemma degree_power_le: "degree (p ^ n) \ degree p * n" by (induct n) (auto intro: order_trans degree_mult_le) lemma coeff_0_power: "coeff (p ^ n) 0 = coeff p 0 ^ n" by (induct n) (simp_all add: coeff_mult) lemma poly_smult [simp]: "poly (smult a p) x = a * poly p x" by (induct p) (simp_all add: algebra_simps) lemma poly_mult [simp]: "poly (p * q) x = poly p x * poly q x" by (induct p) (simp_all add: algebra_simps) lemma poly_power [simp]: "poly (p ^ n) x = poly p x ^ n" for p :: "'a::comm_semiring_1 poly" by (induct n) simp_all lemma poly_prod: "poly (\k\A. p k) x = (\k\A. poly (p k) x)" by (induct A rule: infinite_finite_induct) simp_all lemma degree_prod_sum_le: "finite S \ degree (prod f S) \ sum (degree \ f) S" proof (induct S rule: finite_induct) case empty then show ?case by simp next case (insert a S) show ?case unfolding prod.insert[OF insert(1-2)] sum.insert[OF insert(1-2)] by (rule le_trans[OF degree_mult_le]) (use insert in auto) qed lemma coeff_0_prod_list: "coeff (prod_list xs) 0 = prod_list (map (\p. coeff p 0) xs)" by (induct xs) (simp_all add: coeff_mult) lemma coeff_monom_mult: "coeff (monom c n * p) k = (if k < n then 0 else c * coeff p (k - n))" proof - have "coeff (monom c n * p) k = (\i\k. (if n = i then c else 0) * coeff p (k - i))" by (simp add: coeff_mult) also have "\ = (\i\k. (if n = i then c * coeff p (k - i) else 0))" by (intro sum.cong) simp_all also have "\ = (if k < n then 0 else c * coeff p (k - n))" by simp finally show ?thesis . qed lemma monom_1_dvd_iff': "monom 1 n dvd p \ (\kkkk. coeff p (k + n))" have "\\<^sub>\k. coeff p (k + n) = 0" by (subst cofinite_eq_sequentially, subst eventually_sequentially_seg, subst cofinite_eq_sequentially [symmetric]) transfer then have coeff_r [simp]: "coeff r k = coeff p (k + n)" for k unfolding r_def by (subst poly.Abs_poly_inverse) simp_all have "p = monom 1 n * r" by (rule poly_eqI, subst coeff_monom_mult) (simp_all add: zero) then show "monom 1 n dvd p" by simp qed subsection \Mapping polynomials\ definition map_poly :: "('a :: zero \ 'b :: zero) \ 'a poly \ 'b poly" where "map_poly f p = Poly (map f (coeffs p))" lemma map_poly_0 [simp]: "map_poly f 0 = 0" by (simp add: map_poly_def) lemma map_poly_1: "map_poly f 1 = [:f 1:]" by (simp add: map_poly_def) lemma map_poly_1' [simp]: "f 1 = 1 \ map_poly f 1 = 1" by (simp add: map_poly_def one_pCons) lemma coeff_map_poly: assumes "f 0 = 0" shows "coeff (map_poly f p) n = f (coeff p n)" by (auto simp: assms map_poly_def nth_default_def coeffs_def not_less Suc_le_eq coeff_eq_0 simp del: upt_Suc) lemma coeffs_map_poly [code abstract]: "coeffs (map_poly f p) = strip_while ((=) 0) (map f (coeffs p))" by (simp add: map_poly_def) lemma coeffs_map_poly': assumes "\x. x \ 0 \ f x \ 0" shows "coeffs (map_poly f p) = map f (coeffs p)" using assms by (auto simp add: coeffs_map_poly strip_while_idem_iff last_coeffs_eq_coeff_degree no_trailing_unfold last_map) lemma set_coeffs_map_poly: "(\x. f x = 0 \ x = 0) \ set (coeffs (map_poly f p)) = f ` set (coeffs p)" by (simp add: coeffs_map_poly') lemma degree_map_poly: assumes "\x. x \ 0 \ f x \ 0" shows "degree (map_poly f p) = degree p" by (simp add: degree_eq_length_coeffs coeffs_map_poly' assms) lemma map_poly_eq_0_iff: assumes "f 0 = 0" "\x. x \ set (coeffs p) \ x \ 0 \ f x \ 0" shows "map_poly f p = 0 \ p = 0" proof - have "(coeff (map_poly f p) n = 0) = (coeff p n = 0)" for n proof - have "coeff (map_poly f p) n = f (coeff p n)" by (simp add: coeff_map_poly assms) also have "\ = 0 \ coeff p n = 0" proof (cases "n < length (coeffs p)") case True then have "coeff p n \ set (coeffs p)" by (auto simp: coeffs_def simp del: upt_Suc) with assms show "f (coeff p n) = 0 \ coeff p n = 0" by auto next case False then show ?thesis by (auto simp: assms length_coeffs nth_default_coeffs_eq [symmetric] nth_default_def) qed finally show ?thesis . qed then show ?thesis by (auto simp: poly_eq_iff) qed lemma map_poly_smult: assumes "f 0 = 0""\c x. f (c * x) = f c * f x" shows "map_poly f (smult c p) = smult (f c) (map_poly f p)" by (intro poly_eqI) (simp_all add: assms coeff_map_poly) lemma map_poly_pCons: assumes "f 0 = 0" shows "map_poly f (pCons c p) = pCons (f c) (map_poly f p)" by (intro poly_eqI) (simp_all add: assms coeff_map_poly coeff_pCons split: nat.splits) lemma map_poly_map_poly: assumes "f 0 = 0" "g 0 = 0" shows "map_poly f (map_poly g p) = map_poly (f \ g) p" by (intro poly_eqI) (simp add: coeff_map_poly assms) lemma map_poly_id [simp]: "map_poly id p = p" by (simp add: map_poly_def) lemma map_poly_id' [simp]: "map_poly (\x. x) p = p" by (simp add: map_poly_def) lemma map_poly_cong: assumes "(\x. x \ set (coeffs p) \ f x = g x)" shows "map_poly f p = map_poly g p" proof - from assms have "map f (coeffs p) = map g (coeffs p)" by (intro map_cong) simp_all then show ?thesis by (simp only: coeffs_eq_iff coeffs_map_poly) qed lemma map_poly_monom: "f 0 = 0 \ map_poly f (monom c n) = monom (f c) n" by (intro poly_eqI) (simp_all add: coeff_map_poly) lemma map_poly_idI: assumes "\x. x \ set (coeffs p) \ f x = x" shows "map_poly f p = p" using map_poly_cong[OF assms, of _ id] by simp lemma map_poly_idI': assumes "\x. x \ set (coeffs p) \ f x = x" shows "p = map_poly f p" using map_poly_cong[OF assms, of _ id] by simp lemma smult_conv_map_poly: "smult c p = map_poly (\x. c * x) p" by (intro poly_eqI) (simp_all add: coeff_map_poly) subsection \Conversions\ lemma of_nat_poly: "of_nat n = [:of_nat n:]" by (induct n) (simp_all add: one_pCons) lemma of_nat_monom: "of_nat n = monom (of_nat n) 0" by (simp add: of_nat_poly monom_0) lemma degree_of_nat [simp]: "degree (of_nat n) = 0" by (simp add: of_nat_poly) lemma lead_coeff_of_nat [simp]: "lead_coeff (of_nat n) = of_nat n" by (simp add: of_nat_poly) lemma of_int_poly: "of_int k = [:of_int k:]" by (simp only: of_int_of_nat of_nat_poly) simp lemma of_int_monom: "of_int k = monom (of_int k) 0" by (simp add: of_int_poly monom_0) lemma degree_of_int [simp]: "degree (of_int k) = 0" by (simp add: of_int_poly) lemma lead_coeff_of_int [simp]: "lead_coeff (of_int k) = of_int k" by (simp add: of_int_poly) lemma numeral_poly: "numeral n = [:numeral n:]" proof - have "numeral n = of_nat (numeral n)" by simp also have "\ = [:of_nat (numeral n):]" by (simp add: of_nat_poly) finally show ?thesis by simp qed lemma numeral_monom: "numeral n = monom (numeral n) 0" by (simp add: numeral_poly monom_0) lemma degree_numeral [simp]: "degree (numeral n) = 0" by (simp add: numeral_poly) lemma lead_coeff_numeral [simp]: "lead_coeff (numeral n) = numeral n" by (simp add: numeral_poly) subsection \Lemmas about divisibility\ lemma dvd_smult: assumes "p dvd q" shows "p dvd smult a q" proof - from assms obtain k where "q = p * k" .. then have "smult a q = p * smult a k" by simp then show "p dvd smult a q" .. qed lemma dvd_smult_cancel: "p dvd smult a q \ a \ 0 \ p dvd q" for a :: "'a::field" by (drule dvd_smult [where a="inverse a"]) simp lemma dvd_smult_iff: "a \ 0 \ p dvd smult a q \ p dvd q" for a :: "'a::field" by (safe elim!: dvd_smult dvd_smult_cancel) lemma smult_dvd_cancel: assumes "smult a p dvd q" shows "p dvd q" proof - from assms obtain k where "q = smult a p * k" .. then have "q = p * smult a k" by simp then show "p dvd q" .. qed lemma smult_dvd: "p dvd q \ a \ 0 \ smult a p dvd q" for a :: "'a::field" by (rule smult_dvd_cancel [where a="inverse a"]) simp lemma smult_dvd_iff: "smult a p dvd q \ (if a = 0 then q = 0 else p dvd q)" for a :: "'a::field" by (auto elim: smult_dvd smult_dvd_cancel) lemma is_unit_smult_iff: "smult c p dvd 1 \ c dvd 1 \ p dvd 1" proof - have "smult c p = [:c:] * p" by simp also have "\ dvd 1 \ c dvd 1 \ p dvd 1" proof safe assume *: "[:c:] * p dvd 1" then show "p dvd 1" by (rule dvd_mult_right) from * obtain q where q: "1 = [:c:] * p * q" by (rule dvdE) have "c dvd c * (coeff p 0 * coeff q 0)" by simp also have "\ = coeff ([:c:] * p * q) 0" by (simp add: mult.assoc coeff_mult) also note q [symmetric] finally have "c dvd coeff 1 0" . then show "c dvd 1" by simp next assume "c dvd 1" "p dvd 1" from this(1) obtain d where "1 = c * d" by (rule dvdE) then have "1 = [:c:] * [:d:]" by (simp add: one_pCons ac_simps) then have "[:c:] dvd 1" by (rule dvdI) from mult_dvd_mono[OF this \p dvd 1\] show "[:c:] * p dvd 1" by simp qed finally show ?thesis . qed subsection \Polynomials form an integral domain\ instance poly :: (idom) idom .. instance poly :: ("{ring_char_0, comm_ring_1}") ring_char_0 by standard (auto simp add: of_nat_poly intro: injI) lemma degree_mult_eq: "p \ 0 \ q \ 0 \ degree (p * q) = degree p + degree q" for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" by (rule order_antisym [OF degree_mult_le le_degree]) (simp add: coeff_mult_degree_sum) +lemma degree_prod_eq_sum_degree: + fixes A :: "'a set" + and f :: "'a \ 'b::idom poly" + assumes f0: "\i\A. f i \ 0" + shows "degree (\i\A. (f i)) = (\i\A. degree (f i))" + using assms + by (induction A rule: infinite_finite_induct) (auto simp: degree_mult_eq) + lemma degree_mult_eq_0: "degree (p * q) = 0 \ p = 0 \ q = 0 \ (p \ 0 \ q \ 0 \ degree p = 0 \ degree q = 0)" for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" by (auto simp: degree_mult_eq) lemma degree_power_eq: "p \ 0 \ degree ((p :: 'a :: idom poly) ^ n) = n * degree p" by (induction n) (simp_all add: degree_mult_eq) lemma degree_mult_right_le: fixes p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" assumes "q \ 0" shows "degree p \ degree (p * q)" using assms by (cases "p = 0") (simp_all add: degree_mult_eq) lemma coeff_degree_mult: "coeff (p * q) (degree (p * q)) = coeff q (degree q) * coeff p (degree p)" for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" by (cases "p = 0 \ q = 0") (auto simp: degree_mult_eq coeff_mult_degree_sum mult_ac) lemma dvd_imp_degree_le: "p dvd q \ q \ 0 \ degree p \ degree q" for p q :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" by (erule dvdE, hypsubst, subst degree_mult_eq) auto lemma divides_degree: fixes p q :: "'a ::{comm_semiring_1,semiring_no_zero_divisors} poly" assumes "p dvd q" shows "degree p \ degree q \ q = 0" by (metis dvd_imp_degree_le assms) lemma const_poly_dvd_iff: fixes c :: "'a::{comm_semiring_1,semiring_no_zero_divisors}" shows "[:c:] dvd p \ (\n. c dvd coeff p n)" proof (cases "c = 0 \ p = 0") case True then show ?thesis by (auto intro!: poly_eqI) next case False show ?thesis proof assume "[:c:] dvd p" then show "\n. c dvd coeff p n" by (auto elim!: dvdE simp: coeffs_def) next assume *: "\n. c dvd coeff p n" define mydiv where "mydiv x y = (SOME z. x = y * z)" for x y :: 'a have mydiv: "x = y * mydiv x y" if "y dvd x" for x y using that unfolding mydiv_def dvd_def by (rule someI_ex) define q where "q = Poly (map (\a. mydiv a c) (coeffs p))" from False * have "p = q * [:c:]" by (intro poly_eqI) (auto simp: q_def nth_default_def not_less length_coeffs_degree coeffs_nth intro!: coeff_eq_0 mydiv) then show "[:c:] dvd p" by (simp only: dvd_triv_right) qed qed lemma const_poly_dvd_const_poly_iff [simp]: "[:a:] dvd [:b:] \ a dvd b" for a b :: "'a::{comm_semiring_1,semiring_no_zero_divisors}" by (subst const_poly_dvd_iff) (auto simp: coeff_pCons split: nat.splits) lemma lead_coeff_mult: "lead_coeff (p * q) = lead_coeff p * lead_coeff q" for p q :: "'a::{comm_semiring_0, semiring_no_zero_divisors} poly" by (cases "p = 0 \ q = 0") (auto simp: coeff_mult_degree_sum degree_mult_eq) +lemma lead_coeff_prod: "lead_coeff (prod f A) = (\x\A. lead_coeff (f x))" + for f :: "'a \ 'b::{comm_semiring_1, semiring_no_zero_divisors} poly" + by (induction A rule: infinite_finite_induct) (auto simp: lead_coeff_mult) + lemma lead_coeff_smult: "lead_coeff (smult c p) = c * lead_coeff p" for p :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" proof - have "smult c p = [:c:] * p" by simp also have "lead_coeff \ = c * lead_coeff p" by (subst lead_coeff_mult) simp_all finally show ?thesis . qed lemma lead_coeff_1 [simp]: "lead_coeff 1 = 1" by simp lemma lead_coeff_power: "lead_coeff (p ^ n) = lead_coeff p ^ n" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" by (induct n) (simp_all add: lead_coeff_mult) subsection \Polynomials form an ordered integral domain\ definition pos_poly :: "'a::linordered_semidom poly \ bool" where "pos_poly p \ 0 < coeff p (degree p)" lemma pos_poly_pCons: "pos_poly (pCons a p) \ pos_poly p \ (p = 0 \ 0 < a)" by (simp add: pos_poly_def) lemma not_pos_poly_0 [simp]: "\ pos_poly 0" by (simp add: pos_poly_def) lemma pos_poly_add: "pos_poly p \ pos_poly q \ pos_poly (p + q)" proof (induction p arbitrary: q) case (pCons a p) then show ?case by (cases q; force simp add: pos_poly_pCons add_pos_pos) qed auto lemma pos_poly_mult: "pos_poly p \ pos_poly q \ pos_poly (p * q)" by (simp add: pos_poly_def coeff_degree_mult) lemma pos_poly_total: "p = 0 \ pos_poly p \ pos_poly (- p)" for p :: "'a::linordered_idom poly" by (induct p) (auto simp: pos_poly_pCons) lemma pos_poly_coeffs [code]: "pos_poly p \ (let as = coeffs p in as \ [] \ last as > 0)" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (auto simp add: pos_poly_def last_coeffs_eq_coeff_degree) next assume ?lhs then have *: "0 < coeff p (degree p)" by (simp add: pos_poly_def) then have "p \ 0" by auto with * show ?rhs by (simp add: last_coeffs_eq_coeff_degree) qed instantiation poly :: (linordered_idom) linordered_idom begin definition "x < y \ pos_poly (y - x)" definition "x \ y \ x = y \ pos_poly (y - x)" definition "\x::'a poly\ = (if x < 0 then - x else x)" definition "sgn (x::'a poly) = (if x = 0 then 0 else if 0 < x then 1 else - 1)" instance proof fix x y z :: "'a poly" show "x < y \ x \ y \ \ y \ x" unfolding less_eq_poly_def less_poly_def using pos_poly_add by force then show "x \ y \ y \ x \ x = y" using less_eq_poly_def less_poly_def by force show "x \ x" by (simp add: less_eq_poly_def) show "x \ y \ y \ z \ x \ z" using less_eq_poly_def pos_poly_add by fastforce show "x \ y \ z + x \ z + y" by (simp add: less_eq_poly_def) show "x \ y \ y \ x" unfolding less_eq_poly_def using pos_poly_total [of "x - y"] by auto show "x < y \ 0 < z \ z * x < z * y" by (simp add: less_poly_def right_diff_distrib [symmetric] pos_poly_mult) show "\x\ = (if x < 0 then - x else x)" by (rule abs_poly_def) show "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)" by (rule sgn_poly_def) qed end text \TODO: Simplification rules for comparisons\ subsection \Synthetic division and polynomial roots\ subsubsection \Synthetic division\ text \Synthetic division is simply division by the linear polynomial \<^term>\x - c\.\ definition synthetic_divmod :: "'a::comm_semiring_0 poly \ 'a \ 'a poly \ 'a" where "synthetic_divmod p c = fold_coeffs (\a (q, r). (pCons r q, a + c * r)) p (0, 0)" definition synthetic_div :: "'a::comm_semiring_0 poly \ 'a \ 'a poly" where "synthetic_div p c = fst (synthetic_divmod p c)" lemma synthetic_divmod_0 [simp]: "synthetic_divmod 0 c = (0, 0)" by (simp add: synthetic_divmod_def) lemma synthetic_divmod_pCons [simp]: "synthetic_divmod (pCons a p) c = (\(q, r). (pCons r q, a + c * r)) (synthetic_divmod p c)" by (cases "p = 0 \ a = 0") (auto simp add: synthetic_divmod_def) lemma synthetic_div_0 [simp]: "synthetic_div 0 c = 0" by (simp add: synthetic_div_def) lemma synthetic_div_unique_lemma: "smult c p = pCons a p \ p = 0" by (induct p arbitrary: a) simp_all lemma snd_synthetic_divmod: "snd (synthetic_divmod p c) = poly p c" by (induct p) (simp_all add: split_def) lemma synthetic_div_pCons [simp]: "synthetic_div (pCons a p) c = pCons (poly p c) (synthetic_div p c)" by (simp add: synthetic_div_def split_def snd_synthetic_divmod) lemma synthetic_div_eq_0_iff: "synthetic_div p c = 0 \ degree p = 0" proof (induct p) case 0 then show ?case by simp next case (pCons a p) then show ?case by (cases p) simp qed lemma degree_synthetic_div: "degree (synthetic_div p c) = degree p - 1" by (induct p) (simp_all add: synthetic_div_eq_0_iff) lemma synthetic_div_correct: "p + smult c (synthetic_div p c) = pCons (poly p c) (synthetic_div p c)" by (induct p) simp_all lemma synthetic_div_unique: "p + smult c q = pCons r q \ r = poly p c \ q = synthetic_div p c" proof (induction p arbitrary: q r) case 0 then show ?case using synthetic_div_unique_lemma by fastforce next case (pCons a p) then show ?case by (cases q; force) qed lemma synthetic_div_correct': "[:-c, 1:] * synthetic_div p c + [:poly p c:] = p" for c :: "'a::comm_ring_1" using synthetic_div_correct [of p c] by (simp add: algebra_simps) subsubsection \Polynomial roots\ lemma poly_eq_0_iff_dvd: "poly p c = 0 \ [:- c, 1:] dvd p" (is "?lhs \ ?rhs") for c :: "'a::comm_ring_1" proof assume ?lhs with synthetic_div_correct' [of c p] have "p = [:-c, 1:] * synthetic_div p c" by simp then show ?rhs .. next assume ?rhs then obtain k where "p = [:-c, 1:] * k" by (rule dvdE) then show ?lhs by simp qed lemma dvd_iff_poly_eq_0: "[:c, 1:] dvd p \ poly p (- c) = 0" for c :: "'a::comm_ring_1" by (simp add: poly_eq_0_iff_dvd) lemma poly_roots_finite: "p \ 0 \ finite {x. poly p x = 0}" for p :: "'a::{comm_ring_1,ring_no_zero_divisors} poly" proof (induct n \ "degree p" arbitrary: p) case 0 then obtain a where "a \ 0" and "p = [:a:]" by (cases p) (simp split: if_splits) then show "finite {x. poly p x = 0}" by simp next case (Suc n) show "finite {x. poly p x = 0}" proof (cases "\x. poly p x = 0") case False then show "finite {x. poly p x = 0}" by simp next case True then obtain a where "poly p a = 0" .. then have "[:-a, 1:] dvd p" by (simp only: poly_eq_0_iff_dvd) then obtain k where k: "p = [:-a, 1:] * k" .. with \p \ 0\ have "k \ 0" by auto with k have "degree p = Suc (degree k)" by (simp add: degree_mult_eq del: mult_pCons_left) with \Suc n = degree p\ have "n = degree k" by simp from this \k \ 0\ have "finite {x. poly k x = 0}" by (rule Suc.hyps) then have "finite (insert a {x. poly k x = 0})" by simp then show "finite {x. poly p x = 0}" by (simp add: k Collect_disj_eq del: mult_pCons_left) qed qed lemma poly_eq_poly_eq_iff: "poly p = poly q \ p = q" (is "?lhs \ ?rhs") for p q :: "'a::{comm_ring_1,ring_no_zero_divisors,ring_char_0} poly" proof assume ?rhs then show ?lhs by simp next assume ?lhs have "poly p = poly 0 \ p = 0" for p :: "'a poly" proof (cases "p = 0") case False then show ?thesis by (auto simp add: infinite_UNIV_char_0 dest: poly_roots_finite) qed auto from \?lhs\ and this [of "p - q"] show ?rhs by auto qed lemma poly_all_0_iff_0: "(\x. poly p x = 0) \ p = 0" for p :: "'a::{ring_char_0,comm_ring_1,ring_no_zero_divisors} poly" by (auto simp add: poly_eq_poly_eq_iff [symmetric]) subsubsection \Order of polynomial roots\ definition order :: "'a::idom \ 'a poly \ nat" where "order a p = (LEAST n. \ [:-a, 1:] ^ Suc n dvd p)" lemma coeff_linear_power: "coeff ([:a, 1:] ^ n) n = 1" for a :: "'a::comm_semiring_1" proof (induct n) case (Suc n) have "degree ([:a, 1:] ^ n) \ 1 * n" by (metis One_nat_def degree_pCons_eq_if degree_power_le one_neq_zero one_pCons) then have "coeff ([:a, 1:] ^ n) (Suc n) = 0" by (simp add: coeff_eq_0) then show ?case using Suc.hyps by fastforce qed auto lemma degree_linear_power: "degree ([:a, 1:] ^ n) = n" for a :: "'a::comm_semiring_1" proof (rule order_antisym) show "degree ([:a, 1:] ^ n) \ n" by (metis One_nat_def degree_pCons_eq_if degree_power_le mult.left_neutral one_neq_zero one_pCons) qed (simp add: coeff_linear_power le_degree) lemma order_1: "[:-a, 1:] ^ order a p dvd p" proof (cases "p = 0") case False show ?thesis proof (cases "order a p") case (Suc n) then show ?thesis by (metis lessI not_less_Least order_def) qed auto qed auto lemma order_2: assumes "p \ 0" shows "\ [:-a, 1:] ^ Suc (order a p) dvd p" proof - have False if "[:- a, 1:] ^ Suc (degree p) dvd p" using dvd_imp_degree_le [OF that] by (metis Suc_n_not_le_n assms degree_linear_power) then show ?thesis unfolding order_def by (metis (no_types, lifting) LeastI) qed lemma order: "p \ 0 \ [:-a, 1:] ^ order a p dvd p \ \ [:-a, 1:] ^ Suc (order a p) dvd p" by (rule conjI [OF order_1 order_2]) lemma order_degree: assumes p: "p \ 0" shows "order a p \ degree p" proof - have "order a p = degree ([:-a, 1:] ^ order a p)" by (simp only: degree_linear_power) also from order_1 p have "\ \ degree p" by (rule dvd_imp_degree_le) finally show ?thesis . qed lemma order_root: "poly p a = 0 \ p = 0 \ order a p \ 0" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (metis One_nat_def order_2 poly_eq_0_iff_dvd power_one_right) show "?rhs \ ?lhs" by (meson dvd_power dvd_trans neq0_conv order_1 poly_0 poly_eq_0_iff_dvd) qed lemma order_0I: "poly p a \ 0 \ order a p = 0" by (subst (asm) order_root) auto lemma order_unique_lemma: fixes p :: "'a::idom poly" assumes "[:-a, 1:] ^ n dvd p" "\ [:-a, 1:] ^ Suc n dvd p" shows "order a p = n" unfolding Polynomial.order_def by (metis (mono_tags, lifting) Least_equality assms not_less_eq_eq power_le_dvd) lemma order_mult: assumes "p * q \ 0" shows "order a (p * q) = order a p + order a q" proof - define i where "i \ order a p" define j where "j \ order a q" define t where "t \ [:-a, 1:]" have t_dvd_iff: "\u. t dvd u \ poly u a = 0" by (simp add: t_def dvd_iff_poly_eq_0) have dvd: "t ^ i dvd p" "t ^ j dvd q" and "\ t ^ Suc i dvd p" "\ t ^ Suc j dvd q" using assms i_def j_def order_1 order_2 t_def by auto then have "\ t ^ Suc(i + j) dvd p * q" by (elim dvdE) (simp add: power_add t_dvd_iff) moreover have "t ^ (i + j) dvd p * q" using dvd by (simp add: mult_dvd_mono power_add) ultimately show "order a (p * q) = i + j" using order_unique_lemma t_def by blast qed lemma order_smult: assumes "c \ 0" shows "order x (smult c p) = order x p" proof (cases "p = 0") case True then show ?thesis by simp next case False have "smult c p = [:c:] * p" by simp also from assms False have "order x \ = order x [:c:] + order x p" by (subst order_mult) simp_all also have "order x [:c:] = 0" by (rule order_0I) (use assms in auto) finally show ?thesis by simp qed text \Next three lemmas contributed by Wenda Li\ lemma order_1_eq_0 [simp]:"order x 1 = 0" by (metis order_root poly_1 zero_neq_one) lemma order_uminus[simp]: "order x (-p) = order x p" by (metis neg_equal_0_iff_equal order_smult smult_1_left smult_minus_left) lemma order_power_n_n: "order a ([:-a,1:]^n)=n" proof (induct n) (*might be proved more concisely using nat_less_induct*) case 0 then show ?case by (metis order_root poly_1 power_0 zero_neq_one) next case (Suc n) have "order a ([:- a, 1:] ^ Suc n) = order a ([:- a, 1:] ^ n) + order a [:-a,1:]" by (metis (no_types, hide_lams) One_nat_def add_Suc_right monoid_add_class.add.right_neutral one_neq_zero order_mult pCons_eq_0_iff power_add power_eq_0_iff power_one_right) moreover have "order a [:-a,1:] = 1" unfolding order_def proof (rule Least_equality, rule notI) assume "[:- a, 1:] ^ Suc 1 dvd [:- a, 1:]" then have "degree ([:- a, 1:] ^ Suc 1) \ degree ([:- a, 1:])" by (rule dvd_imp_degree_le) auto then show False by auto next fix y assume *: "\ [:- a, 1:] ^ Suc y dvd [:- a, 1:]" show "1 \ y" proof (rule ccontr) assume "\ 1 \ y" then have "y = 0" by auto then have "[:- a, 1:] ^ Suc y dvd [:- a, 1:]" by auto with * show False by auto qed qed ultimately show ?case using Suc by auto qed lemma order_0_monom [simp]: "c \ 0 \ order 0 (monom c n) = n" using order_power_n_n[of 0 n] by (simp add: monom_altdef order_smult) lemma dvd_imp_order_le: "q \ 0 \ p dvd q \ Polynomial.order a p \ Polynomial.order a q" by (auto simp: order_mult elim: dvdE) text \Now justify the standard squarefree decomposition, i.e. \f / gcd f f'\.\ lemma order_divides: "[:-a, 1:] ^ n dvd p \ p = 0 \ n \ order a p" by (meson dvd_0_right not_less_eq_eq order_1 order_2 power_le_dvd) lemma order_decomp: assumes "p \ 0" shows "\q. p = [:- a, 1:] ^ order a p * q \ \ [:- a, 1:] dvd q" proof - from assms have *: "[:- a, 1:] ^ order a p dvd p" and **: "\ [:- a, 1:] ^ Suc (order a p) dvd p" by (auto dest: order) from * obtain q where q: "p = [:- a, 1:] ^ order a p * q" .. with ** have "\ [:- a, 1:] ^ Suc (order a p) dvd [:- a, 1:] ^ order a p * q" by simp then have "\ [:- a, 1:] ^ order a p * [:- a, 1:] dvd [:- a, 1:] ^ order a p * q" by simp with idom_class.dvd_mult_cancel_left [of "[:- a, 1:] ^ order a p" "[:- a, 1:]" q] have "\ [:- a, 1:] dvd q" by auto with q show ?thesis by blast qed lemma monom_1_dvd_iff: "p \ 0 \ monom 1 n dvd p \ n \ order 0 p" using order_divides[of 0 n p] by (simp add: monom_altdef) subsection \Additional induction rules on polynomials\ text \ An induction rule for induction over the roots of a polynomial with a certain property. (e.g. all positive roots) \ lemma poly_root_induct [case_names 0 no_roots root]: fixes p :: "'a :: idom poly" assumes "Q 0" and "\p. (\a. P a \ poly p a \ 0) \ Q p" and "\a p. P a \ Q p \ Q ([:a, -1:] * p)" shows "Q p" proof (induction "degree p" arbitrary: p rule: less_induct) case (less p) show ?case proof (cases "p = 0") case True with assms(1) show ?thesis by simp next case False show ?thesis proof (cases "\a. P a \ poly p a = 0") case False then show ?thesis by (intro assms(2)) blast next case True then obtain a where a: "P a" "poly p a = 0" by blast then have "-[:-a, 1:] dvd p" by (subst minus_dvd_iff) (simp add: poly_eq_0_iff_dvd) then obtain q where q: "p = [:a, -1:] * q" by (elim dvdE) simp with False have "q \ 0" by auto have "degree p = Suc (degree q)" by (subst q, subst degree_mult_eq) (simp_all add: \q \ 0\) then have "Q q" by (intro less) simp with a(1) have "Q ([:a, -1:] * q)" by (rule assms(3)) with q show ?thesis by simp qed qed qed lemma dropWhile_replicate_append: "dropWhile ((=) a) (replicate n a @ ys) = dropWhile ((=) a) ys" by (induct n) simp_all lemma Poly_append_replicate_0: "Poly (xs @ replicate n 0) = Poly xs" by (subst coeffs_eq_iff) (simp_all add: strip_while_def dropWhile_replicate_append) text \ An induction rule for simultaneous induction over two polynomials, prepending one coefficient in each step. \ lemma poly_induct2 [case_names 0 pCons]: assumes "P 0 0" "\a p b q. P p q \ P (pCons a p) (pCons b q)" shows "P p q" proof - define n where "n = max (length (coeffs p)) (length (coeffs q))" define xs where "xs = coeffs p @ (replicate (n - length (coeffs p)) 0)" define ys where "ys = coeffs q @ (replicate (n - length (coeffs q)) 0)" have "length xs = length ys" by (simp add: xs_def ys_def n_def) then have "P (Poly xs) (Poly ys)" by (induct rule: list_induct2) (simp_all add: assms) also have "Poly xs = p" by (simp add: xs_def Poly_append_replicate_0) also have "Poly ys = q" by (simp add: ys_def Poly_append_replicate_0) finally show ?thesis . qed subsection \Composition of polynomials\ (* Several lemmas contributed by René Thiemann and Akihisa Yamada *) definition pcompose :: "'a::comm_semiring_0 poly \ 'a poly \ 'a poly" where "pcompose p q = fold_coeffs (\a c. [:a:] + q * c) p 0" notation pcompose (infixl "\\<^sub>p" 71) lemma pcompose_0 [simp]: "pcompose 0 q = 0" by (simp add: pcompose_def) lemma pcompose_pCons: "pcompose (pCons a p) q = [:a:] + q * pcompose p q" by (cases "p = 0 \ a = 0") (auto simp add: pcompose_def) lemma pcompose_1: "pcompose 1 p = 1" for p :: "'a::comm_semiring_1 poly" by (auto simp: one_pCons pcompose_pCons) lemma poly_pcompose: "poly (pcompose p q) x = poly p (poly q x)" by (induct p) (simp_all add: pcompose_pCons) lemma degree_pcompose_le: "degree (pcompose p q) \ degree p * degree q" proof (induction p) case (pCons a p) then show ?case proof (clarsimp simp add: pcompose_pCons) assume "degree (p \\<^sub>p q) \ degree p * degree q" "p \ 0" then have "degree (q * p \\<^sub>p q) \ degree q + degree p * degree q" by (meson add_le_cancel_left degree_mult_le dual_order.trans pCons.IH) then show "degree ([:a:] + q * p \\<^sub>p q) \ degree q + degree p * degree q" by (simp add: degree_add_le) qed qed auto lemma pcompose_add: "pcompose (p + q) r = pcompose p r + pcompose q r" for p q r :: "'a::{comm_semiring_0, ab_semigroup_add} poly" proof (induction p q rule: poly_induct2) case 0 then show ?case by simp next case (pCons a p b q) have "pcompose (pCons a p + pCons b q) r = [:a + b:] + r * pcompose p r + r * pcompose q r" by (simp_all add: pcompose_pCons pCons.IH algebra_simps) also have "[:a + b:] = [:a:] + [:b:]" by simp also have "\ + r * pcompose p r + r * pcompose q r = pcompose (pCons a p) r + pcompose (pCons b q) r" by (simp only: pcompose_pCons add_ac) finally show ?case . qed lemma pcompose_uminus: "pcompose (-p) r = -pcompose p r" for p r :: "'a::comm_ring poly" by (induct p) (simp_all add: pcompose_pCons) lemma pcompose_diff: "pcompose (p - q) r = pcompose p r - pcompose q r" for p q r :: "'a::comm_ring poly" using pcompose_add[of p "-q"] by (simp add: pcompose_uminus) lemma pcompose_smult: "pcompose (smult a p) r = smult a (pcompose p r)" for p r :: "'a::comm_semiring_0 poly" by (induct p) (simp_all add: pcompose_pCons pcompose_add smult_add_right) lemma pcompose_mult: "pcompose (p * q) r = pcompose p r * pcompose q r" for p q r :: "'a::comm_semiring_0 poly" by (induct p arbitrary: q) (simp_all add: pcompose_add pcompose_smult pcompose_pCons algebra_simps) lemma pcompose_assoc: "pcompose p (pcompose q r) = pcompose (pcompose p q) r" for p q r :: "'a::comm_semiring_0 poly" by (induct p arbitrary: q) (simp_all add: pcompose_pCons pcompose_add pcompose_mult) lemma pcompose_idR[simp]: "pcompose p [: 0, 1 :] = p" for p :: "'a::comm_semiring_1 poly" by (induct p) (simp_all add: pcompose_pCons) lemma pcompose_sum: "pcompose (sum f A) p = sum (\i. pcompose (f i) p) A" by (induct A rule: infinite_finite_induct) (simp_all add: pcompose_1 pcompose_add) lemma pcompose_prod: "pcompose (prod f A) p = prod (\i. pcompose (f i) p) A" by (induct A rule: infinite_finite_induct) (simp_all add: pcompose_1 pcompose_mult) lemma pcompose_const [simp]: "pcompose [:a:] q = [:a:]" by (subst pcompose_pCons) simp lemma pcompose_0': "pcompose p 0 = [:coeff p 0:]" by (induct p) (auto simp add: pcompose_pCons) lemma degree_pcompose: "degree (pcompose p q) = degree p * degree q" for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" proof (induct p) case 0 then show ?case by auto next case (pCons a p) consider "degree (q * pcompose p q) = 0" | "degree (q * pcompose p q) > 0" by blast then show ?case proof cases case prems: 1 show ?thesis proof (cases "p = 0") case True then show ?thesis by auto next case False from prems have "degree q = 0 \ pcompose p q = 0" by (auto simp add: degree_mult_eq_0) moreover have False if "pcompose p q = 0" "degree q \ 0" proof - from pCons.hyps(2) that have "degree p = 0" by auto then obtain a1 where "p = [:a1:]" by (metis degree_pCons_eq_if old.nat.distinct(2) pCons_cases) with \pcompose p q = 0\ \p \ 0\ show False by auto qed ultimately have "degree (pCons a p) * degree q = 0" by auto moreover have "degree (pcompose (pCons a p) q) = 0" proof - from prems have "0 = max (degree [:a:]) (degree (q * pcompose p q))" by simp also have "\ \ degree ([:a:] + q * pcompose p q)" by (rule degree_add_le_max) finally show ?thesis by (auto simp add: pcompose_pCons) qed ultimately show ?thesis by simp qed next case prems: 2 then have "p \ 0" "q \ 0" "pcompose p q \ 0" by auto from prems degree_add_eq_right [of "[:a:]"] have "degree (pcompose (pCons a p) q) = degree (q * pcompose p q)" by (auto simp: pcompose_pCons) with pCons.hyps(2) degree_mult_eq[OF \q\0\ \pcompose p q\0\] show ?thesis by auto qed qed lemma pcompose_eq_0: fixes p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" assumes "pcompose p q = 0" "degree q > 0" shows "p = 0" proof - from assms degree_pcompose [of p q] have "degree p = 0" by auto then obtain a where "p = [:a:]" by (metis degree_pCons_eq_if gr0_conv_Suc neq0_conv pCons_cases) with assms(1) have "a = 0" by auto with \p = [:a:]\ show ?thesis by simp qed lemma lead_coeff_comp: fixes p q :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" assumes "degree q > 0" shows "lead_coeff (pcompose p q) = lead_coeff p * lead_coeff q ^ (degree p)" proof (induct p) case 0 then show ?case by auto next case (pCons a p) consider "degree (q * pcompose p q) = 0" | "degree (q * pcompose p q) > 0" by blast then show ?case proof cases case prems: 1 then have "pcompose p q = 0" by (metis assms degree_0 degree_mult_eq_0 neq0_conv) with pcompose_eq_0[OF _ \degree q > 0\] have "p = 0" by simp then show ?thesis by auto next case prems: 2 then have "degree [:a:] < degree (q * pcompose p q)" by simp then have "lead_coeff ([:a:] + q * p \\<^sub>p q) = lead_coeff (q * p \\<^sub>p q)" by (rule lead_coeff_add_le) then have "lead_coeff (pcompose (pCons a p) q) = lead_coeff (q * pcompose p q)" by (simp add: pcompose_pCons) also have "\ = lead_coeff q * (lead_coeff p * lead_coeff q ^ degree p)" using pCons.hyps(2) lead_coeff_mult[of q "pcompose p q"] by simp also have "\ = lead_coeff p * lead_coeff q ^ (degree p + 1)" by (auto simp: mult_ac) finally show ?thesis by auto qed qed +subsection \Closure properties of coefficients\ + + +context + fixes R :: "'a :: comm_semiring_1 set" + assumes R_0: "0 \ R" + assumes R_plus: "\x y. x \ R \ y \ R \ x + y \ R" + assumes R_mult: "\x y. x \ R \ y \ R \ x * y \ R" +begin + +lemma coeff_mult_semiring_closed: + assumes "\i. coeff p i \ R" "\i. coeff q i \ R" + shows "coeff (p * q) i \ R" +proof - + have R_sum: "sum f A \ R" if "\x. x \ A \ f x \ R" for A and f :: "nat \ 'a" + using that by (induction A rule: infinite_finite_induct) (auto intro: R_0 R_plus) + show ?thesis + unfolding coeff_mult by (auto intro!: R_sum R_mult assms) +qed + +lemma coeff_pcompose_semiring_closed: + assumes "\i. coeff p i \ R" "\i. coeff q i \ R" + shows "coeff (pcompose p q) i \ R" + using assms(1) +proof (induction p arbitrary: i) + case (pCons a p i) + have [simp]: "a \ R" + using pCons.prems[of 0] by auto + have "coeff p i \ R" for i + using pCons.prems[of "Suc i"] by auto + hence "coeff (p \\<^sub>p q) i \ R" for i + using pCons.prems by (intro pCons.IH) + thus ?case + by (auto simp: pcompose_pCons coeff_pCons split: nat.splits + intro!: assms R_plus coeff_mult_semiring_closed) +qed auto + +end + + subsection \Shifting polynomials\ definition poly_shift :: "nat \ 'a::zero poly \ 'a poly" where "poly_shift n p = Abs_poly (\i. coeff p (i + n))" lemma nth_default_drop: "nth_default x (drop n xs) m = nth_default x xs (m + n)" by (auto simp add: nth_default_def add_ac) lemma nth_default_take: "nth_default x (take n xs) m = (if m < n then nth_default x xs m else x)" by (auto simp add: nth_default_def add_ac) lemma coeff_poly_shift: "coeff (poly_shift n p) i = coeff p (i + n)" proof - from MOST_coeff_eq_0[of p] obtain m where "\k>m. coeff p k = 0" by (auto simp: MOST_nat) then have "\k>m. coeff p (k + n) = 0" by auto then have "\\<^sub>\k. coeff p (k + n) = 0" by (auto simp: MOST_nat) then show ?thesis by (simp add: poly_shift_def poly.Abs_poly_inverse) qed lemma poly_shift_id [simp]: "poly_shift 0 = (\x. x)" by (simp add: poly_eq_iff fun_eq_iff coeff_poly_shift) lemma poly_shift_0 [simp]: "poly_shift n 0 = 0" by (simp add: poly_eq_iff coeff_poly_shift) lemma poly_shift_1: "poly_shift n 1 = (if n = 0 then 1 else 0)" by (simp add: poly_eq_iff coeff_poly_shift) lemma poly_shift_monom: "poly_shift n (monom c m) = (if m \ n then monom c (m - n) else 0)" by (auto simp add: poly_eq_iff coeff_poly_shift) lemma coeffs_shift_poly [code abstract]: "coeffs (poly_shift n p) = drop n (coeffs p)" proof (cases "p = 0") case True then show ?thesis by simp next case False then show ?thesis by (intro coeffs_eqI) (simp_all add: coeff_poly_shift nth_default_drop nth_default_coeffs_eq) qed subsection \Truncating polynomials\ definition poly_cutoff where "poly_cutoff n p = Abs_poly (\k. if k < n then coeff p k else 0)" lemma coeff_poly_cutoff: "coeff (poly_cutoff n p) k = (if k < n then coeff p k else 0)" unfolding poly_cutoff_def by (subst poly.Abs_poly_inverse) (auto simp: MOST_nat intro: exI[of _ n]) lemma poly_cutoff_0 [simp]: "poly_cutoff n 0 = 0" by (simp add: poly_eq_iff coeff_poly_cutoff) lemma poly_cutoff_1 [simp]: "poly_cutoff n 1 = (if n = 0 then 0 else 1)" by (simp add: poly_eq_iff coeff_poly_cutoff) lemma coeffs_poly_cutoff [code abstract]: "coeffs (poly_cutoff n p) = strip_while ((=) 0) (take n (coeffs p))" proof (cases "strip_while ((=) 0) (take n (coeffs p)) = []") case True then have "coeff (poly_cutoff n p) k = 0" for k unfolding coeff_poly_cutoff by (auto simp: nth_default_coeffs_eq [symmetric] nth_default_def set_conv_nth) then have "poly_cutoff n p = 0" by (simp add: poly_eq_iff) then show ?thesis by (subst True) simp_all next case False have "no_trailing ((=) 0) (strip_while ((=) 0) (take n (coeffs p)))" by simp with False have "last (strip_while ((=) 0) (take n (coeffs p))) \ 0" unfolding no_trailing_unfold by auto then show ?thesis by (intro coeffs_eqI) (simp_all add: coeff_poly_cutoff nth_default_take nth_default_coeffs_eq) qed subsection \Reflecting polynomials\ definition reflect_poly :: "'a::zero poly \ 'a poly" where "reflect_poly p = Poly (rev (coeffs p))" lemma coeffs_reflect_poly [code abstract]: "coeffs (reflect_poly p) = rev (dropWhile ((=) 0) (coeffs p))" by (simp add: reflect_poly_def) lemma reflect_poly_0 [simp]: "reflect_poly 0 = 0" by (simp add: reflect_poly_def) lemma reflect_poly_1 [simp]: "reflect_poly 1 = 1" by (simp add: reflect_poly_def one_pCons) lemma coeff_reflect_poly: "coeff (reflect_poly p) n = (if n > degree p then 0 else coeff p (degree p - n))" by (cases "p = 0") (auto simp add: reflect_poly_def nth_default_def rev_nth degree_eq_length_coeffs coeffs_nth not_less dest: le_imp_less_Suc) lemma coeff_0_reflect_poly_0_iff [simp]: "coeff (reflect_poly p) 0 = 0 \ p = 0" by (simp add: coeff_reflect_poly) lemma reflect_poly_at_0_eq_0_iff [simp]: "poly (reflect_poly p) 0 = 0 \ p = 0" by (simp add: coeff_reflect_poly poly_0_coeff_0) lemma reflect_poly_pCons': "p \ 0 \ reflect_poly (pCons c p) = reflect_poly p + monom c (Suc (degree p))" by (intro poly_eqI) (auto simp: coeff_reflect_poly coeff_pCons not_less Suc_diff_le split: nat.split) lemma reflect_poly_const [simp]: "reflect_poly [:a:] = [:a:]" by (cases "a = 0") (simp_all add: reflect_poly_def) lemma poly_reflect_poly_nz: "x \ 0 \ poly (reflect_poly p) x = x ^ degree p * poly p (inverse x)" for x :: "'a::field" by (induct rule: pCons_induct) (simp_all add: field_simps reflect_poly_pCons' poly_monom) lemma coeff_0_reflect_poly [simp]: "coeff (reflect_poly p) 0 = lead_coeff p" by (simp add: coeff_reflect_poly) lemma poly_reflect_poly_0 [simp]: "poly (reflect_poly p) 0 = lead_coeff p" by (simp add: poly_0_coeff_0) lemma reflect_poly_reflect_poly [simp]: "coeff p 0 \ 0 \ reflect_poly (reflect_poly p) = p" by (cases p rule: pCons_cases) (simp add: reflect_poly_def ) lemma degree_reflect_poly_le: "degree (reflect_poly p) \ degree p" by (simp add: degree_eq_length_coeffs coeffs_reflect_poly length_dropWhile_le diff_le_mono) lemma reflect_poly_pCons: "a \ 0 \ reflect_poly (pCons a p) = Poly (rev (a # coeffs p))" by (subst coeffs_eq_iff) (simp add: coeffs_reflect_poly) lemma degree_reflect_poly_eq [simp]: "coeff p 0 \ 0 \ degree (reflect_poly p) = degree p" by (cases p rule: pCons_cases) (simp add: reflect_poly_pCons degree_eq_length_coeffs) (* TODO: does this work with zero divisors as well? Probably not. *) lemma reflect_poly_mult: "reflect_poly (p * q) = reflect_poly p * reflect_poly q" for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" proof (cases "p = 0 \ q = 0") case False then have [simp]: "p \ 0" "q \ 0" by auto show ?thesis proof (rule poly_eqI) show "coeff (reflect_poly (p * q)) i = coeff (reflect_poly p * reflect_poly q) i" for i proof (cases "i \ degree (p * q)") case True define A where "A = {..i} \ {i - degree q..degree p}" define B where "B = {..degree p} \ {degree p - i..degree (p*q) - i}" let ?f = "\j. degree p - j" from True have "coeff (reflect_poly (p * q)) i = coeff (p * q) (degree (p * q) - i)" by (simp add: coeff_reflect_poly) also have "\ = (\j\degree (p * q) - i. coeff p j * coeff q (degree (p * q) - i - j))" by (simp add: coeff_mult) also have "\ = (\j\B. coeff p j * coeff q (degree (p * q) - i - j))" by (intro sum.mono_neutral_right) (auto simp: B_def degree_mult_eq not_le coeff_eq_0) also from True have "\ = (\j\A. coeff p (degree p - j) * coeff q (degree q - (i - j)))" by (intro sum.reindex_bij_witness[of _ ?f ?f]) (auto simp: A_def B_def degree_mult_eq add_ac) also have "\ = (\j\i. if j \ {i - degree q..degree p} then coeff p (degree p - j) * coeff q (degree q - (i - j)) else 0)" by (subst sum.inter_restrict [symmetric]) (simp_all add: A_def) also have "\ = coeff (reflect_poly p * reflect_poly q) i" by (fastforce simp: coeff_mult coeff_reflect_poly intro!: sum.cong) finally show ?thesis . qed (auto simp: coeff_mult coeff_reflect_poly coeff_eq_0 degree_mult_eq intro!: sum.neutral) qed qed auto lemma reflect_poly_smult: "reflect_poly (smult c p) = smult c (reflect_poly p)" for p :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" using reflect_poly_mult[of "[:c:]" p] by simp lemma reflect_poly_power: "reflect_poly (p ^ n) = reflect_poly p ^ n" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" by (induct n) (simp_all add: reflect_poly_mult) lemma reflect_poly_prod: "reflect_poly (prod f A) = prod (\x. reflect_poly (f x)) A" for f :: "_ \ _::{comm_semiring_0,semiring_no_zero_divisors} poly" by (induct A rule: infinite_finite_induct) (simp_all add: reflect_poly_mult) lemma reflect_poly_prod_list: "reflect_poly (prod_list xs) = prod_list (map reflect_poly xs)" for xs :: "_::{comm_semiring_0,semiring_no_zero_divisors} poly list" by (induct xs) (simp_all add: reflect_poly_mult) lemma reflect_poly_Poly_nz: "no_trailing (HOL.eq 0) xs \ reflect_poly (Poly xs) = Poly (rev xs)" by (simp add: reflect_poly_def) lemmas reflect_poly_simps = reflect_poly_0 reflect_poly_1 reflect_poly_const reflect_poly_smult reflect_poly_mult reflect_poly_power reflect_poly_prod reflect_poly_prod_list subsection \Derivatives\ function pderiv :: "('a :: {comm_semiring_1,semiring_no_zero_divisors}) poly \ 'a poly" where "pderiv (pCons a p) = (if p = 0 then 0 else p + pCons 0 (pderiv p))" by (auto intro: pCons_cases) termination pderiv by (relation "measure degree") simp_all declare pderiv.simps[simp del] lemma pderiv_0 [simp]: "pderiv 0 = 0" using pderiv.simps [of 0 0] by simp lemma pderiv_pCons: "pderiv (pCons a p) = p + pCons 0 (pderiv p)" by (simp add: pderiv.simps) lemma pderiv_1 [simp]: "pderiv 1 = 0" by (simp add: one_pCons pderiv_pCons) lemma pderiv_of_nat [simp]: "pderiv (of_nat n) = 0" and pderiv_numeral [simp]: "pderiv (numeral m) = 0" by (simp_all add: of_nat_poly numeral_poly pderiv_pCons) lemma coeff_pderiv: "coeff (pderiv p) n = of_nat (Suc n) * coeff p (Suc n)" by (induct p arbitrary: n) (auto simp add: pderiv_pCons coeff_pCons algebra_simps split: nat.split) fun pderiv_coeffs_code :: "'a::{comm_semiring_1,semiring_no_zero_divisors} \ 'a list \ 'a list" where "pderiv_coeffs_code f (x # xs) = cCons (f * x) (pderiv_coeffs_code (f+1) xs)" | "pderiv_coeffs_code f [] = []" definition pderiv_coeffs :: "'a::{comm_semiring_1,semiring_no_zero_divisors} list \ 'a list" where "pderiv_coeffs xs = pderiv_coeffs_code 1 (tl xs)" (* Efficient code for pderiv contributed by René Thiemann and Akihisa Yamada *) lemma pderiv_coeffs_code: "nth_default 0 (pderiv_coeffs_code f xs) n = (f + of_nat n) * nth_default 0 xs n" proof (induct xs arbitrary: f n) case Nil then show ?case by simp next case (Cons x xs) show ?case proof (cases n) case 0 then show ?thesis by (cases "pderiv_coeffs_code (f + 1) xs = [] \ f * x = 0") (auto simp: cCons_def) next case n: (Suc m) show ?thesis proof (cases "pderiv_coeffs_code (f + 1) xs = [] \ f * x = 0") case False then have "nth_default 0 (pderiv_coeffs_code f (x # xs)) n = nth_default 0 (pderiv_coeffs_code (f + 1) xs) m" by (auto simp: cCons_def n) also have "\ = (f + of_nat n) * nth_default 0 xs m" by (simp add: Cons n add_ac) finally show ?thesis by (simp add: n) next case True have empty: "pderiv_coeffs_code g xs = [] \ g + of_nat m = 0 \ nth_default 0 xs m = 0" for g proof (induct xs arbitrary: g m) case Nil then show ?case by simp next case (Cons x xs) from Cons(2) have empty: "pderiv_coeffs_code (g + 1) xs = []" and g: "g = 0 \ x = 0" by (auto simp: cCons_def split: if_splits) note IH = Cons(1)[OF empty] from IH[of m] IH[of "m - 1"] g show ?case by (cases m) (auto simp: field_simps) qed from True have "nth_default 0 (pderiv_coeffs_code f (x # xs)) n = 0" by (auto simp: cCons_def n) moreover from True have "(f + of_nat n) * nth_default 0 (x # xs) n = 0" by (simp add: n) (use empty[of "f+1"] in \auto simp: field_simps\) ultimately show ?thesis by simp qed qed qed lemma coeffs_pderiv_code [code abstract]: "coeffs (pderiv p) = pderiv_coeffs (coeffs p)" unfolding pderiv_coeffs_def proof (rule coeffs_eqI, unfold pderiv_coeffs_code coeff_pderiv, goal_cases) case (1 n) have id: "coeff p (Suc n) = nth_default 0 (map (\i. coeff p (Suc i)) [0.. degree p = 0" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" proof (cases "degree p") case 0 then show ?thesis by (metis degree_eq_zeroE pderiv.simps) next case (Suc n) then show ?thesis - by (metis coeff_0 coeff_pderiv degree_0 leading_coeff_0_iff mult_eq_0_iff nat.distinct(1) of_nat_eq_0_iff) + using coeff_0 coeff_pderiv degree_0 leading_coeff_0_iff mult_eq_0_iff nat.distinct(1) of_nat_eq_0_iff + by (metis coeff_0 coeff_pderiv degree_0 leading_coeff_0_iff mult_eq_0_iff nat.distinct(1) of_nat_eq_0_iff) qed lemma degree_pderiv: "degree (pderiv p) = degree p - 1" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" proof - have "degree p - 1 \ degree (pderiv p)" proof (cases "degree p") case (Suc n) then show ?thesis by (metis coeff_pderiv degree_0 diff_Suc_1 le_degree leading_coeff_0_iff mult_eq_0_iff nat.distinct(1) of_nat_eq_0_iff) qed auto moreover have "\i>degree p - 1. coeff (pderiv p) i = 0" by (simp add: coeff_eq_0 coeff_pderiv) ultimately show ?thesis using order_antisym [OF degree_le] by blast qed lemma not_dvd_pderiv: fixes p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" assumes "degree p \ 0" shows "\ p dvd pderiv p" proof assume dvd: "p dvd pderiv p" then obtain q where p: "pderiv p = p * q" unfolding dvd_def by auto from dvd have le: "degree p \ degree (pderiv p)" by (simp add: assms dvd_imp_degree_le pderiv_eq_0_iff) from assms and this [unfolded degree_pderiv] show False by auto qed lemma dvd_pderiv_iff [simp]: "p dvd pderiv p \ degree p = 0" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" using not_dvd_pderiv[of p] by (auto simp: pderiv_eq_0_iff [symmetric]) lemma pderiv_singleton [simp]: "pderiv [:a:] = 0" by (simp add: pderiv_pCons) lemma pderiv_add: "pderiv (p + q) = pderiv p + pderiv q" by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps) lemma pderiv_minus: "pderiv (- p :: 'a :: idom poly) = - pderiv p" by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps) lemma pderiv_diff: "pderiv ((p :: _ :: idom poly) - q) = pderiv p - pderiv q" by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps) lemma pderiv_smult: "pderiv (smult a p) = smult a (pderiv p)" by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps) lemma pderiv_mult: "pderiv (p * q) = p * pderiv q + q * pderiv p" by (induct p) (auto simp: pderiv_add pderiv_smult pderiv_pCons algebra_simps) lemma pderiv_power_Suc: "pderiv (p ^ Suc n) = smult (of_nat (Suc n)) (p ^ n) * pderiv p" proof (induction n) case (Suc n) then show ?case by (simp add: pderiv_mult smult_add_left algebra_simps) qed auto lemma pderiv_pcompose: "pderiv (pcompose p q) = pcompose (pderiv p) q * pderiv q" by (induction p rule: pCons_induct) (auto simp: pcompose_pCons pderiv_add pderiv_mult pderiv_pCons pcompose_add algebra_simps) lemma pderiv_prod: "pderiv (prod f (as)) = (\a\as. prod f (as - {a}) * pderiv (f a))" proof (induct as rule: infinite_finite_induct) case (insert a as) then have id: "prod f (insert a as) = f a * prod f as" "\g. sum g (insert a as) = g a + sum g as" "insert a as - {a} = as" by auto have "prod f (insert a as - {b}) = f a * prod f (as - {b})" if "b \ as" for b proof - from \a \ as\ that have *: "insert a as - {b} = insert a (as - {b})" by auto show ?thesis unfolding * by (subst prod.insert) (use insert in auto) qed then show ?case unfolding id pderiv_mult insert(3) sum_distrib_left by (auto simp add: ac_simps intro!: sum.cong) qed auto lemma DERIV_pow2: "DERIV (\x. x ^ Suc n) x :> real (Suc n) * (x ^ n)" by (rule DERIV_cong, rule DERIV_pow) simp declare DERIV_pow2 [simp] DERIV_pow [simp] lemma DERIV_add_const: "DERIV f x :> D \ DERIV (\x. a + f x :: 'a::real_normed_field) x :> D" by (rule DERIV_cong, rule DERIV_add) auto lemma poly_DERIV [simp]: "DERIV (\x. poly p x) x :> poly (pderiv p) x" by (induct p) (auto intro!: derivative_eq_intros simp add: pderiv_pCons) lemma poly_isCont[simp]: fixes x::"'a::real_normed_field" shows "isCont (\x. poly p x) x" by (rule poly_DERIV [THEN DERIV_isCont]) lemma tendsto_poly [tendsto_intros]: "(f \ a) F \ ((\x. poly p (f x)) \ poly p a) F" for f :: "_ \ 'a::real_normed_field" by (rule isCont_tendsto_compose [OF poly_isCont]) lemma continuous_within_poly: "continuous (at z within s) (poly p)" for z :: "'a::{real_normed_field}" by (simp add: continuous_within tendsto_poly) lemma continuous_poly [continuous_intros]: "continuous F f \ continuous F (\x. poly p (f x))" for f :: "_ \ 'a::real_normed_field" unfolding continuous_def by (rule tendsto_poly) lemma continuous_on_poly [continuous_intros]: fixes p :: "'a :: {real_normed_field} poly" assumes "continuous_on A f" shows "continuous_on A (\x. poly p (f x))" by (metis DERIV_continuous_on assms continuous_on_compose2 poly_DERIV subset_UNIV) text \Consequences of the derivative theorem above.\ lemma poly_differentiable[simp]: "(\x. poly p x) differentiable (at x)" for x :: real by (simp add: real_differentiable_def) (blast intro: poly_DERIV) lemma poly_IVT_pos: "a < b \ poly p a < 0 \ 0 < poly p b \ \x. a < x \ x < b \ poly p x = 0" for a b :: real using IVT [of "poly p" a 0 b] by (auto simp add: order_le_less) lemma poly_IVT_neg: "a < b \ 0 < poly p a \ poly p b < 0 \ \x. a < x \ x < b \ poly p x = 0" for a b :: real using poly_IVT_pos [where p = "- p"] by simp lemma poly_IVT: "a < b \ poly p a * poly p b < 0 \ \x>a. x < b \ poly p x = 0" for p :: "real poly" by (metis less_not_sym mult_less_0_iff poly_IVT_neg poly_IVT_pos) lemma poly_MVT: "a < b \ \x. a < x \ x < b \ poly p b - poly p a = (b - a) * poly (pderiv p) x" for a b :: real by (simp add: MVT2) lemma poly_MVT': fixes a b :: real assumes "{min a b..max a b} \ A" shows "\x\A. poly p b - poly p a = (b - a) * poly (pderiv p) x" proof (cases a b rule: linorder_cases) case less from poly_MVT[OF less, of p] guess x by (elim exE conjE) then show ?thesis by (intro bexI[of _ x]) (auto intro!: subsetD[OF assms]) next case greater from poly_MVT[OF greater, of p] guess x by (elim exE conjE) then show ?thesis by (intro bexI[of _ x]) (auto simp: algebra_simps intro!: subsetD[OF assms]) qed (use assms in auto) lemma poly_pinfty_gt_lc: fixes p :: "real poly" assumes "lead_coeff p > 0" shows "\n. \ x \ n. poly p x \ lead_coeff p" using assms proof (induct p) case 0 then show ?case by auto next case (pCons a p) from this(1) consider "a \ 0" "p = 0" | "p \ 0" by auto then show ?case proof cases case 1 then show ?thesis by auto next case 2 with pCons obtain n1 where gte_lcoeff: "\x\n1. lead_coeff p \ poly p x" by auto from pCons(3) \p \ 0\ have gt_0: "lead_coeff p > 0" by auto define n where "n = max n1 (1 + \a\ / lead_coeff p)" have "lead_coeff (pCons a p) \ poly (pCons a p) x" if "n \ x" for x proof - from gte_lcoeff that have "lead_coeff p \ poly p x" by (auto simp: n_def) with gt_0 have "\a\ / lead_coeff p \ \a\ / poly p x" and "poly p x > 0" by (auto intro: frac_le) with \n \ x\[unfolded n_def] have "x \ 1 + \a\ / poly p x" by auto with \lead_coeff p \ poly p x\ \poly p x > 0\ \p \ 0\ show "lead_coeff (pCons a p) \ poly (pCons a p) x" by (auto simp: field_simps) qed then show ?thesis by blast qed qed lemma lemma_order_pderiv1: "pderiv ([:- a, 1:] ^ Suc n * q) = [:- a, 1:] ^ Suc n * pderiv q + smult (of_nat (Suc n)) (q * [:- a, 1:] ^ n)" by (simp only: pderiv_mult pderiv_power_Suc) (simp del: power_Suc of_nat_Suc add: pderiv_pCons) lemma lemma_order_pderiv: fixes p :: "'a :: field_char_0 poly" assumes n: "0 < n" and pd: "pderiv p \ 0" and pe: "p = [:- a, 1:] ^ n * q" and nd: "\ [:- a, 1:] dvd q" shows "n = Suc (order a (pderiv p))" proof - from assms have "pderiv ([:- a, 1:] ^ n * q) \ 0" by auto from assms obtain n' where "n = Suc n'" "0 < Suc n'" "pderiv ([:- a, 1:] ^ Suc n' * q) \ 0" by (cases n) auto have "order a (pderiv ([:- a, 1:] ^ Suc n' * q)) = n'" proof (rule order_unique_lemma) show "[:- a, 1:] ^ n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" unfolding lemma_order_pderiv1 proof (rule dvd_add) show "[:- a, 1:] ^ n' dvd [:- a, 1:] ^ Suc n' * pderiv q" by (metis dvdI dvd_mult2 power_Suc2) show "[:- a, 1:] ^ n' dvd smult (of_nat (Suc n')) (q * [:- a, 1:] ^ n')" by (metis dvd_smult dvd_triv_right) qed have "k dvd k * pderiv q + smult (of_nat (Suc n')) l \ k dvd l" for k l by (auto simp del: of_nat_Suc simp: dvd_add_right_iff dvd_smult_iff) then show "\ [:- a, 1:] ^ Suc n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" unfolding lemma_order_pderiv1 by (metis nd dvd_mult_cancel_right power_not_zero pCons_eq_0_iff power_Suc zero_neq_one) qed then show ?thesis by (metis \n = Suc n'\ pe) qed lemma order_pderiv: "order a p = Suc (order a (pderiv p))" if "pderiv p \ 0" "order a p \ 0" for p :: "'a::field_char_0 poly" proof (cases "p = 0") case False obtain q where "p = [:- a, 1:] ^ order a p * q \ \ [:- a, 1:] dvd q" using False order_decomp by blast then show ?thesis using lemma_order_pderiv that by blast qed (use that in auto) lemma poly_squarefree_decomp_order: fixes p :: "'a::field_char_0 poly" assumes "pderiv p \ 0" and p: "p = q * d" and p': "pderiv p = e * d" and d: "d = r * p + s * pderiv p" shows "order a q = (if order a p = 0 then 0 else 1)" proof (rule classical) assume 1: "\ ?thesis" from \pderiv p \ 0\ have "p \ 0" by auto with p have "order a p = order a q + order a d" by (simp add: order_mult) with 1 have "order a p \ 0" by (auto split: if_splits) from \pderiv p \ 0\ \pderiv p = e * d\ have oapp: "order a (pderiv p) = order a e + order a d" by (simp add: order_mult) from \pderiv p \ 0\ \order a p \ 0\ have oap: "order a p = Suc (order a (pderiv p))" by (rule order_pderiv) from \p \ 0\ \p = q * d\ have "d \ 0" by simp have "[:- a, 1:] ^ order a (pderiv p) dvd r * p" by (metis dvd_trans dvd_triv_right oap order_1 power_Suc) then have "([:-a, 1:] ^ (order a (pderiv p))) dvd d" by (simp add: d order_1) with \d \ 0\ have "order a (pderiv p) \ order a d" by (simp add: order_divides) show ?thesis using \order a p = order a q + order a d\ and oapp oap and \order a (pderiv p) \ order a d\ by auto qed lemma poly_squarefree_decomp_order2: "pderiv p \ 0 \ p = q * d \ pderiv p = e * d \ d = r * p + s * pderiv p \ \a. order a q = (if order a p = 0 then 0 else 1)" for p :: "'a::field_char_0 poly" by (blast intro: poly_squarefree_decomp_order) lemma order_pderiv2: "pderiv p \ 0 \ order a p \ 0 \ order a (pderiv p) = n \ order a p = Suc n" for p :: "'a::field_char_0 poly" by (auto dest: order_pderiv) definition rsquarefree :: "'a::idom poly \ bool" where "rsquarefree p \ p \ 0 \ (\a. order a p = 0 \ order a p = 1)" lemma pderiv_iszero: "pderiv p = 0 \ \h. p = [:h:]" for p :: "'a::{semidom,semiring_char_0} poly" by (cases p) (auto simp: pderiv_eq_0_iff split: if_splits) lemma rsquarefree_roots: "rsquarefree p \ (\a. \ (poly p a = 0 \ poly (pderiv p) a = 0))" for p :: "'a::field_char_0 poly" proof (cases "p = 0") case False show ?thesis proof (cases "pderiv p = 0") case True with \p \ 0\ pderiv_iszero show ?thesis by (force simp add: order_0I rsquarefree_def) next case False with \p \ 0\ order_pderiv2 show ?thesis by (force simp add: rsquarefree_def order_root) qed qed (simp add: rsquarefree_def) lemma poly_squarefree_decomp: fixes p :: "'a::field_char_0 poly" assumes "pderiv p \ 0" and "p = q * d" and "pderiv p = e * d" and "d = r * p + s * pderiv p" shows "rsquarefree q \ (\a. poly q a = 0 \ poly p a = 0)" proof - from \pderiv p \ 0\ have "p \ 0" by auto with \p = q * d\ have "q \ 0" by simp from assms have "\a. order a q = (if order a p = 0 then 0 else 1)" by (rule poly_squarefree_decomp_order2) with \p \ 0\ \q \ 0\ show ?thesis by (simp add: rsquarefree_def order_root) qed subsection \Algebraic numbers\ text \ Algebraic numbers can be defined in two equivalent ways: all real numbers that are roots of rational polynomials or of integer polynomials. The Algebraic-Numbers AFP entry uses the rational definition, but we need the integer definition. The equivalence is obvious since any rational polynomial can be multiplied with the LCM of its coefficients, yielding an integer polynomial with the same roots. \ definition algebraic :: "'a :: field_char_0 \ bool" where "algebraic x \ (\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0)" lemma algebraicI: "(\i. coeff p i \ \) \ p \ 0 \ poly p x = 0 \ algebraic x" unfolding algebraic_def by blast lemma algebraicE: assumes "algebraic x" obtains p where "\i. coeff p i \ \" "p \ 0" "poly p x = 0" using assms unfolding algebraic_def by blast lemma algebraic_altdef: "algebraic x \ (\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0)" for p :: "'a::field_char_0 poly" proof safe fix p assume rat: "\i. coeff p i \ \" and root: "poly p x = 0" and nz: "p \ 0" define cs where "cs = coeffs p" from rat have "\c\range (coeff p). \c'. c = of_rat c'" unfolding Rats_def by blast then obtain f where f: "coeff p i = of_rat (f (coeff p i))" for i by (subst (asm) bchoice_iff) blast define cs' where "cs' = map (quotient_of \ f) (coeffs p)" define d where "d = Lcm (set (map snd cs'))" define p' where "p' = smult (of_int d) p" have "coeff p' n \ \" for n proof (cases "n \ degree p") case True define c where "c = coeff p n" define a where "a = fst (quotient_of (f (coeff p n)))" define b where "b = snd (quotient_of (f (coeff p n)))" have b_pos: "b > 0" unfolding b_def using quotient_of_denom_pos' by simp have "coeff p' n = of_int d * coeff p n" by (simp add: p'_def) also have "coeff p n = of_rat (of_int a / of_int b)" unfolding a_def b_def by (subst quotient_of_div [of "f (coeff p n)", symmetric]) (simp_all add: f [symmetric]) also have "of_int d * \ = of_rat (of_int (a*d) / of_int b)" by (simp add: of_rat_mult of_rat_divide) also from nz True have "b \ snd ` set cs'" by (force simp: cs'_def o_def b_def coeffs_def simp del: upt_Suc) then have "b dvd (a * d)" by (simp add: d_def) then have "of_int (a * d) / of_int b \ (\ :: rat set)" by (rule of_int_divide_in_Ints) then have "of_rat (of_int (a * d) / of_int b) \ \" by (elim Ints_cases) auto finally show ?thesis . next case False then show ?thesis by (auto simp: p'_def not_le coeff_eq_0) qed moreover have "set (map snd cs') \ {0<..}" unfolding cs'_def using quotient_of_denom_pos' by (auto simp: coeffs_def simp del: upt_Suc) then have "d \ 0" unfolding d_def by (induct cs') simp_all with nz have "p' \ 0" by (simp add: p'_def) moreover from root have "poly p' x = 0" by (simp add: p'_def) ultimately show "algebraic x" unfolding algebraic_def by blast next assume "algebraic x" then obtain p where p: "coeff p i \ \" "poly p x = 0" "p \ 0" for i by (force simp: algebraic_def) moreover have "coeff p i \ \ \ coeff p i \ \" for i by (elim Ints_cases) simp ultimately show "\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0" by auto qed +subsection \Algebraic integers\ + +inductive algebraic_int :: "'a :: field \ bool" where + "\lead_coeff p = 1; \i. coeff p i \ \; poly p x = 0\ \ algebraic_int x" + +lemma algebraic_int_altdef_ipoly: + fixes x :: "'a :: field_char_0" + shows "algebraic_int x \ (\p. poly (map_poly of_int p) x = 0 \ lead_coeff p = 1)" +proof + assume "algebraic_int x" + then obtain p where p: "lead_coeff p = 1" "\i. coeff p i \ \" "poly p x = 0" + by (auto elim: algebraic_int.cases) + define the_int where "the_int = (\x::'a. THE r. x = of_int r)" + define p' where "p' = map_poly the_int p" + have of_int_the_int: "of_int (the_int x) = x" if "x \ \" for x + unfolding the_int_def by (rule sym, rule theI') (insert that, auto simp: Ints_def) + have the_int_0_iff: "the_int x = 0 \ x = 0" if "x \ \" for x + using of_int_the_int[OF that] by auto + have [simp]: "the_int 0 = 0" + by (subst the_int_0_iff) auto + have "map_poly of_int p' = map_poly (of_int \ the_int) p" + by (simp add: p'_def map_poly_map_poly) + also from p of_int_the_int have "\ = p" + by (subst poly_eq_iff) (auto simp: coeff_map_poly) + finally have p_p': "map_poly of_int p' = p" . + + show "(\p. poly (map_poly of_int p) x = 0 \ lead_coeff p = 1)" + proof (intro exI conjI notI) + from p show "poly (map_poly of_int p') x = 0" by (simp add: p_p') + next + show "lead_coeff p' = 1" + using p by (simp flip: p_p' add: degree_map_poly coeff_map_poly) + qed +next + assume "\p. poly (map_poly of_int p) x = 0 \ lead_coeff p = 1" + then obtain p where p: "poly (map_poly of_int p) x = 0" "lead_coeff p = 1" + by auto + define p' where "p' = (map_poly of_int p :: 'a poly)" + from p have "lead_coeff p' = 1" "poly p' x = 0" "\i. coeff p' i \ \" + by (auto simp: p'_def coeff_map_poly degree_map_poly) + thus "algebraic_int x" + by (intro algebraic_int.intros) +qed + +theorem rational_algebraic_int_is_int: + assumes "algebraic_int x" and "x \ \" + shows "x \ \" +proof - + from assms(2) obtain a b where ab: "b > 0" "Rings.coprime a b" and x_eq: "x = of_int a / of_int b" + by (auto elim: Rats_cases') + from \b > 0\ have [simp]: "b \ 0" + by auto + from assms(1) obtain p + where p: "lead_coeff p = 1" "\i. coeff p i \ \" "poly p x = 0" + by (auto simp: algebraic_int.simps) + + define q :: "'a poly" where "q = [:-of_int a, of_int b:]" + have "poly q x = 0" "q \ 0" "\i. coeff q i \ \" + by (auto simp: x_eq q_def coeff_pCons split: nat.splits) + define n where "n = degree p" + have "n > 0" + using p by (intro Nat.gr0I) (auto simp: n_def elim!: degree_eq_zeroE) + have "(\i \" + using p by auto + then obtain R where R: "of_int R = (\i = (\i\n. coeff p i * x ^ i * of_int b ^ n)" + by (simp add: poly_altdef n_def sum_distrib_right) + also have "\ = (\i\n. coeff p i * of_int (a ^ i * b ^ (n - i)))" + by (intro sum.cong) (auto simp: x_eq field_simps simp flip: power_add) + also have "{..n} = insert n {..n > 0\ by auto + also have "(\i\\. coeff p i * of_int (a ^ i * b ^ (n - i))) = + coeff p n * of_int (a ^ n) + (\iii = of_int (b * R)" + by (simp add: R sum_distrib_left sum_distrib_right mult_ac) + finally have "of_int (a ^ n) = (-of_int (b * R) :: 'a)" + by (auto simp: add_eq_0_iff) + hence "a ^ n = -b * R" + by (simp flip: of_int_mult of_int_power of_int_minus) + hence "b dvd a ^ n" + by simp + with \Rings.coprime a b\ have "b dvd 1" + by (meson coprime_power_left_iff dvd_refl not_coprimeI) + with x_eq and \b > 0\ show ?thesis + by auto +qed + +lemma algebraic_int_imp_algebraic [dest]: "algebraic_int x \ algebraic x" + by (auto simp: algebraic_int.simps algebraic_def) + +lemma int_imp_algebraic_int: + assumes "x \ \" + shows "algebraic_int x" +proof + show "\i. coeff [:-x, 1:] i \ \" + using assms by (auto simp: coeff_pCons split: nat.splits) +qed auto + +lemma algebraic_int_0 [simp, intro]: "algebraic_int 0" + and algebraic_int_1 [simp, intro]: "algebraic_int 1" + and algebraic_int_numeral [simp, intro]: "algebraic_int (numeral n)" + and algebraic_int_of_nat [simp, intro]: "algebraic_int (of_nat k)" + and algebraic_int_of_int [simp, intro]: "algebraic_int (of_int k)" + by (simp_all add: int_imp_algebraic_int) + +lemma algebraic_int_ii [simp, intro]: "algebraic_int \" +proof + show "poly [:1, 0, 1:] \ = 0" + by simp +qed (auto simp: coeff_pCons split: nat.splits) + +lemma algebraic_int_minus [intro]: + assumes "algebraic_int x" + shows "algebraic_int (-x)" +proof - + from assms obtain p where p: "lead_coeff p = 1" "\i. coeff p i \ \" "poly p x = 0" + by (auto simp: algebraic_int.simps) + define s where "s = (if even (degree p) then 1 else -1 :: 'a)" + + define q where "q = Polynomial.smult s (pcompose p [:0, -1:])" + have "lead_coeff q = s * lead_coeff (pcompose p [:0, -1:])" + by (simp add: q_def) + also have "lead_coeff (pcompose p [:0, -1:]) = lead_coeff p * (- 1) ^ degree p" + by (subst lead_coeff_comp) auto + finally have "poly q (-x) = 0" and "lead_coeff q = 1" + using p by (auto simp: q_def poly_pcompose s_def) + moreover have "coeff q i \ \" for i + proof - + have "coeff (pcompose p [:0, -1:]) i \ \" + using p by (intro coeff_pcompose_semiring_closed) (auto simp: coeff_pCons split: nat.splits) + thus ?thesis by (simp add: q_def s_def) + qed + ultimately show ?thesis + by (auto simp: algebraic_int.simps) +qed + +lemma algebraic_int_minus_iff [simp]: + "algebraic_int (-x) \ algebraic_int (x :: 'a :: field_char_0)" + using algebraic_int_minus[of x] algebraic_int_minus[of "-x"] by auto + +lemma algebraic_int_inverse [intro]: + assumes "poly p x = 0" and "\i. coeff p i \ \" and "coeff p 0 = 1" + shows "algebraic_int (inverse x)" +proof + from assms have [simp]: "x \ 0" + by (auto simp: poly_0_coeff_0) + show "poly (reflect_poly p) (inverse x) = 0" + using assms by (simp add: poly_reflect_poly_nz) +qed (use assms in \auto simp: coeff_reflect_poly\) + +lemma algebraic_int_root: + assumes "algebraic_int y" + and "poly p x = y" and "\i. coeff p i \ \" and "lead_coeff p = 1" and "degree p > 0" + shows "algebraic_int x" +proof - + from assms obtain q where q: "poly q y = 0" "\i. coeff q i \ \" "lead_coeff q = 1" + by (auto simp: algebraic_int.simps) + show ?thesis + proof + from assms q show "lead_coeff (pcompose q p) = 1" + by (subst lead_coeff_comp) auto + from assms q show "\i. coeff (pcompose q p) i \ \" + by (intro allI coeff_pcompose_semiring_closed) auto + show "poly (pcompose q p) x = 0" + using assms q by (simp add: poly_pcompose) + qed +qed + +lemma algebraic_int_abs_real [simp]: + "algebraic_int \x :: real\ \ algebraic_int x" + by (auto simp: abs_if) + +lemma algebraic_int_nth_root_real [intro]: + assumes "algebraic_int x" + shows "algebraic_int (root n x)" +proof (cases "n = 0") + case False + show ?thesis + proof (rule algebraic_int_root) + show "poly (monom 1 n) (root n x) = (if even n then \x\ else x)" + using sgn_power_root[of n x] False + by (auto simp add: poly_monom sgn_if split: if_splits) + qed (use False assms in \auto simp: degree_monom_eq\) +qed auto + +lemma algebraic_int_sqrt [intro]: "algebraic_int x \ algebraic_int (sqrt x)" + by (auto simp: sqrt_def) + +lemma algebraic_int_csqrt [intro]: "algebraic_int x \ algebraic_int (csqrt x)" + by (rule algebraic_int_root[where p = "monom 1 2"]) + (auto simp: poly_monom degree_monom_eq) + +lemma poly_map_poly_cnj [simp]: "poly (map_poly cnj p) x = cnj (poly p (cnj x))" + by (induction p) (auto simp: map_poly_pCons) + +lemma algebraic_int_cnj [intro]: + assumes "algebraic_int x" + shows "algebraic_int (cnj x)" +proof - + from assms obtain p where p: "lead_coeff p = 1" "\i. coeff p i \ \" "poly p x = 0" + by (auto simp: algebraic_int.simps) + show ?thesis + proof + show "poly (map_poly cnj p) (cnj x) = 0" + using p by simp + show "lead_coeff (map_poly cnj p) = 1" + using p by (simp add: coeff_map_poly degree_map_poly) + show "\i. coeff (map_poly cnj p) i \ \" + using p by (auto simp: coeff_map_poly) + qed +qed + +lemma algebraic_int_cnj_iff [simp]: "algebraic_int (cnj x) \ algebraic_int x" + using algebraic_int_cnj[of x] algebraic_int_cnj[of "cnj x"] by auto + +lemma algebraic_int_of_real [intro]: + assumes "algebraic_int x" + shows "algebraic_int (of_real x)" +proof - + from assms obtain p where p: "poly p x = 0" "\i. coeff p i \ \" "lead_coeff p = 1" + by (auto simp: algebraic_int.simps) + show "algebraic_int (of_real x :: 'a)" + proof + have "poly (map_poly of_real p) (of_real x) = (of_real (poly p x) :: 'a)" + by (induction p) (auto simp: map_poly_pCons) + thus "poly (map_poly of_real p) (of_real x) = (0 :: 'a)" + using p by simp + qed (use p in \auto simp: coeff_map_poly degree_map_poly\) +qed + +lemma algebraic_int_of_real_iff [simp]: + "algebraic_int (of_real x :: 'a :: {field_char_0, real_algebra_1}) \ algebraic_int x" +proof + assume "algebraic_int (of_real x :: 'a)" + then obtain p + where p: "poly (map_poly of_int p) (of_real x :: 'a) = 0" "lead_coeff p = 1" + by (auto simp: algebraic_int_altdef_ipoly) + show "algebraic_int x" + unfolding algebraic_int_altdef_ipoly + proof (intro exI[of _ p] conjI) + have "of_real (poly (map_poly real_of_int p) x) = poly (map_poly of_int p) (of_real x :: 'a)" + by (induction p) (auto simp: map_poly_pCons) + also note p(1) + finally show "poly (map_poly real_of_int p) x = 0" by simp + qed (use p in auto) +qed auto + + subsection \Division of polynomials\ subsubsection \Division in general\ instantiation poly :: (idom_divide) idom_divide begin fun divide_poly_main :: "'a \ 'a poly \ 'a poly \ 'a poly \ nat \ nat \ 'a poly" where "divide_poly_main lc q r d dr (Suc n) = (let cr = coeff r dr; a = cr div lc; mon = monom a n in if False \ a * lc = cr then \ \\False \\ is only because of problem in function-package\ divide_poly_main lc (q + mon) (r - mon * d) d (dr - 1) n else 0)" | "divide_poly_main lc q r d dr 0 = q" definition divide_poly :: "'a poly \ 'a poly \ 'a poly" where "divide_poly f g = (if g = 0 then 0 else divide_poly_main (coeff g (degree g)) 0 f g (degree f) (1 + length (coeffs f) - length (coeffs g)))" lemma divide_poly_main: assumes d: "d \ 0" "lc = coeff d (degree d)" and "degree (d * r) \ dr" "divide_poly_main lc q (d * r) d dr n = q'" and "n = 1 + dr - degree d \ dr = 0 \ n = 0 \ d * r = 0" shows "q' = q + r" using assms(3-) proof (induct n arbitrary: q r dr) case (Suc n) let ?rr = "d * r" let ?a = "coeff ?rr dr" let ?qq = "?a div lc" define b where [simp]: "b = monom ?qq n" let ?rrr = "d * (r - b)" let ?qqq = "q + b" note res = Suc(3) from Suc(4) have dr: "dr = n + degree d" by auto from d have lc: "lc \ 0" by auto have "coeff (b * d) dr = coeff b n * coeff d (degree d)" proof (cases "?qq = 0") case True then show ?thesis by simp next case False then have n: "n = degree b" by (simp add: degree_monom_eq) show ?thesis unfolding n dr by (simp add: coeff_mult_degree_sum) qed also have "\ = lc * coeff b n" by (simp add: d) finally have c2: "coeff (b * d) dr = lc * coeff b n" . have rrr: "?rrr = ?rr - b * d" by (simp add: field_simps) have c1: "coeff (d * r) dr = lc * coeff r n" proof (cases "degree r = n") case True with Suc(2) show ?thesis unfolding dr using coeff_mult_degree_sum[of d r] d by (auto simp: ac_simps) next case False from dr Suc(2) have "degree r \ n" by auto (metis add.commute add_le_cancel_left d(1) degree_0 degree_mult_eq diff_is_0_eq diff_zero le_cases) with False have r_n: "degree r < n" by auto then have right: "lc * coeff r n = 0" by (simp add: coeff_eq_0) have "coeff (d * r) dr = coeff (d * r) (degree d + n)" by (simp add: dr ac_simps) also from r_n have "\ = 0" by (metis False Suc.prems(1) add.commute add_left_imp_eq coeff_degree_mult coeff_eq_0 coeff_mult_degree_sum degree_mult_le dr le_eq_less_or_eq) finally show ?thesis by (simp only: right) qed have c0: "coeff ?rrr dr = 0" and id: "lc * (coeff (d * r) dr div lc) = coeff (d * r) dr" unfolding rrr coeff_diff c2 unfolding b_def coeff_monom coeff_smult c1 using lc by auto from res[unfolded divide_poly_main.simps[of lc q] Let_def] id have res: "divide_poly_main lc ?qqq ?rrr d (dr - 1) n = q'" by (simp del: divide_poly_main.simps add: field_simps) note IH = Suc(1)[OF _ res] from Suc(4) have dr: "dr = n + degree d" by auto from Suc(2) have deg_rr: "degree ?rr \ dr" by auto have deg_bd: "degree (b * d) \ dr" unfolding dr b_def by (rule order.trans[OF degree_mult_le]) (auto simp: degree_monom_le) have "degree ?rrr \ dr" unfolding rrr by (rule degree_diff_le[OF deg_rr deg_bd]) with c0 have deg_rrr: "degree ?rrr \ (dr - 1)" by (rule coeff_0_degree_minus_1) have "n = 1 + (dr - 1) - degree d \ dr - 1 = 0 \ n = 0 \ ?rrr = 0" proof (cases dr) case 0 with Suc(4) have 0: "dr = 0" "n = 0" "degree d = 0" by auto with deg_rrr have "degree ?rrr = 0" by simp from degree_eq_zeroE[OF this] obtain a where rrr: "?rrr = [:a:]" by metis show ?thesis unfolding 0 using c0 unfolding rrr 0 by simp next case _: Suc with Suc(4) show ?thesis by auto qed from IH[OF deg_rrr this] show ?case by simp next case 0 show ?case proof (cases "r = 0") case True with 0 show ?thesis by auto next case False from d False have "degree (d * r) = degree d + degree r" by (subst degree_mult_eq) auto with 0 d show ?thesis by auto qed qed lemma divide_poly_main_0: "divide_poly_main 0 0 r d dr n = 0" proof (induct n arbitrary: r d dr) case 0 then show ?case by simp next case Suc show ?case unfolding divide_poly_main.simps[of _ _ r] Let_def by (simp add: Suc del: divide_poly_main.simps) qed lemma divide_poly: assumes g: "g \ 0" shows "(f * g) div g = (f :: 'a poly)" proof - have len: "length (coeffs f) = Suc (degree f)" if "f \ 0" for f :: "'a poly" using that unfolding degree_eq_length_coeffs by auto have "divide_poly_main (coeff g (degree g)) 0 (g * f) g (degree (g * f)) (1 + length (coeffs (g * f)) - length (coeffs g)) = (f * g) div g" by (simp add: divide_poly_def Let_def ac_simps) note main = divide_poly_main[OF g refl le_refl this] have "(f * g) div g = 0 + f" proof (rule main, goal_cases) case 1 show ?case proof (cases "f = 0") case True with g show ?thesis by (auto simp: degree_eq_length_coeffs) next case False with g have fg: "g * f \ 0" by auto show ?thesis unfolding len[OF fg] len[OF g] by auto qed qed then show ?thesis by simp qed lemma divide_poly_0: "f div 0 = 0" for f :: "'a poly" by (simp add: divide_poly_def Let_def divide_poly_main_0) instance by standard (auto simp: divide_poly divide_poly_0) end instance poly :: (idom_divide) algebraic_semidom .. lemma div_const_poly_conv_map_poly: assumes "[:c:] dvd p" shows "p div [:c:] = map_poly (\x. x div c) p" proof (cases "c = 0") case True then show ?thesis by (auto intro!: poly_eqI simp: coeff_map_poly) next case False from assms obtain q where p: "p = [:c:] * q" by (rule dvdE) moreover { have "smult c q = [:c:] * q" by simp also have "\ div [:c:] = q" by (rule nonzero_mult_div_cancel_left) (use False in auto) finally have "smult c q div [:c:] = q" . } ultimately show ?thesis by (intro poly_eqI) (auto simp: coeff_map_poly False) qed lemma is_unit_monom_0: fixes a :: "'a::field" assumes "a \ 0" shows "is_unit (monom a 0)" proof from assms show "1 = monom a 0 * monom (inverse a) 0" by (simp add: mult_monom) qed lemma is_unit_triv: "a \ 0 \ is_unit [:a:]" for a :: "'a::field" by (simp add: is_unit_monom_0 monom_0 [symmetric]) lemma is_unit_iff_degree: fixes p :: "'a::field poly" assumes "p \ 0" shows "is_unit p \ degree p = 0" (is "?lhs \ ?rhs") proof assume ?rhs then obtain a where "p = [:a:]" by (rule degree_eq_zeroE) with assms show ?lhs by (simp add: is_unit_triv) next assume ?lhs then obtain q where "q \ 0" "p * q = 1" .. then have "degree (p * q) = degree 1" by simp with \p \ 0\ \q \ 0\ have "degree p + degree q = 0" by (simp add: degree_mult_eq) then show ?rhs by simp qed lemma is_unit_pCons_iff: "is_unit (pCons a p) \ p = 0 \ a \ 0" for p :: "'a::field poly" by (cases "p = 0") (auto simp: is_unit_triv is_unit_iff_degree) lemma is_unit_monom_trivial: "is_unit p \ monom (coeff p (degree p)) 0 = p" for p :: "'a::field poly" by (cases p) (simp_all add: monom_0 is_unit_pCons_iff) lemma is_unit_const_poly_iff: "[:c:] dvd 1 \ c dvd 1" for c :: "'a::{comm_semiring_1,semiring_no_zero_divisors}" by (auto simp: one_pCons) lemma is_unit_polyE: fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly" assumes "p dvd 1" obtains c where "p = [:c:]" "c dvd 1" proof - from assms obtain q where "1 = p * q" by (rule dvdE) then have "p \ 0" and "q \ 0" by auto from \1 = p * q\ have "degree 1 = degree (p * q)" by simp also from \p \ 0\ and \q \ 0\ have "\ = degree p + degree q" by (simp add: degree_mult_eq) finally have "degree p = 0" by simp with degree_eq_zeroE obtain c where c: "p = [:c:]" . with \p dvd 1\ have "c dvd 1" by (simp add: is_unit_const_poly_iff) with c show thesis .. qed lemma is_unit_polyE': fixes p :: "'a::field poly" assumes "is_unit p" obtains a where "p = monom a 0" and "a \ 0" proof - obtain a q where "p = pCons a q" by (cases p) with assms have "p = [:a:]" and "a \ 0" by (simp_all add: is_unit_pCons_iff) with that show thesis by (simp add: monom_0) qed lemma is_unit_poly_iff: "p dvd 1 \ (\c. p = [:c:] \ c dvd 1)" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" by (auto elim: is_unit_polyE simp add: is_unit_const_poly_iff) subsubsection \Pseudo-Division\ text \This part is by René Thiemann and Akihisa Yamada.\ fun pseudo_divmod_main :: "'a :: comm_ring_1 \ 'a poly \ 'a poly \ 'a poly \ nat \ nat \ 'a poly \ 'a poly" where "pseudo_divmod_main lc q r d dr (Suc n) = (let rr = smult lc r; qq = coeff r dr; rrr = rr - monom qq n * d; qqq = smult lc q + monom qq n in pseudo_divmod_main lc qqq rrr d (dr - 1) n)" | "pseudo_divmod_main lc q r d dr 0 = (q,r)" definition pseudo_divmod :: "'a :: comm_ring_1 poly \ 'a poly \ 'a poly \ 'a poly" where "pseudo_divmod p q \ if q = 0 then (0, p) else pseudo_divmod_main (coeff q (degree q)) 0 p q (degree p) (1 + length (coeffs p) - length (coeffs q))" lemma pseudo_divmod_main: assumes d: "d \ 0" "lc = coeff d (degree d)" and "degree r \ dr" "pseudo_divmod_main lc q r d dr n = (q',r')" and "n = 1 + dr - degree d \ dr = 0 \ n = 0 \ r = 0" shows "(r' = 0 \ degree r' < degree d) \ smult (lc^n) (d * q + r) = d * q' + r'" using assms(3-) proof (induct n arbitrary: q r dr) case 0 then show ?case by auto next case (Suc n) let ?rr = "smult lc r" let ?qq = "coeff r dr" define b where [simp]: "b = monom ?qq n" let ?rrr = "?rr - b * d" let ?qqq = "smult lc q + b" note res = Suc(3) from res[unfolded pseudo_divmod_main.simps[of lc q] Let_def] have res: "pseudo_divmod_main lc ?qqq ?rrr d (dr - 1) n = (q',r')" by (simp del: pseudo_divmod_main.simps) from Suc(4) have dr: "dr = n + degree d" by auto have "coeff (b * d) dr = coeff b n * coeff d (degree d)" proof (cases "?qq = 0") case True then show ?thesis by auto next case False then have n: "n = degree b" by (simp add: degree_monom_eq) show ?thesis unfolding n dr by (simp add: coeff_mult_degree_sum) qed also have "\ = lc * coeff b n" by (simp add: d) finally have "coeff (b * d) dr = lc * coeff b n" . moreover have "coeff ?rr dr = lc * coeff r dr" by simp ultimately have c0: "coeff ?rrr dr = 0" by auto from Suc(4) have dr: "dr = n + degree d" by auto have deg_rr: "degree ?rr \ dr" using Suc(2) degree_smult_le dual_order.trans by blast have deg_bd: "degree (b * d) \ dr" unfolding dr by (rule order.trans[OF degree_mult_le]) (auto simp: degree_monom_le) have "degree ?rrr \ dr" using degree_diff_le[OF deg_rr deg_bd] by auto with c0 have deg_rrr: "degree ?rrr \ (dr - 1)" by (rule coeff_0_degree_minus_1) have "n = 1 + (dr - 1) - degree d \ dr - 1 = 0 \ n = 0 \ ?rrr = 0" proof (cases dr) case 0 with Suc(4) have 0: "dr = 0" "n = 0" "degree d = 0" by auto with deg_rrr have "degree ?rrr = 0" by simp then have "\a. ?rrr = [:a:]" by (metis degree_pCons_eq_if old.nat.distinct(2) pCons_cases) from this obtain a where rrr: "?rrr = [:a:]" by auto show ?thesis unfolding 0 using c0 unfolding rrr 0 by simp next case _: Suc with Suc(4) show ?thesis by auto qed note IH = Suc(1)[OF deg_rrr res this] show ?case proof (intro conjI) from IH show "r' = 0 \ degree r' < degree d" by blast show "smult (lc ^ Suc n) (d * q + r) = d * q' + r'" unfolding IH[THEN conjunct2,symmetric] by (simp add: field_simps smult_add_right) qed qed lemma pseudo_divmod: assumes g: "g \ 0" and *: "pseudo_divmod f g = (q,r)" shows "smult (coeff g (degree g) ^ (Suc (degree f) - degree g)) f = g * q + r" (is ?A) and "r = 0 \ degree r < degree g" (is ?B) proof - from *[unfolded pseudo_divmod_def Let_def] have "pseudo_divmod_main (coeff g (degree g)) 0 f g (degree f) (1 + length (coeffs f) - length (coeffs g)) = (q, r)" by (auto simp: g) note main = pseudo_divmod_main[OF _ _ _ this, OF g refl le_refl] from g have "1 + length (coeffs f) - length (coeffs g) = 1 + degree f - degree g \ degree f = 0 \ 1 + length (coeffs f) - length (coeffs g) = 0 \ f = 0" by (cases "f = 0"; cases "coeffs g") (auto simp: degree_eq_length_coeffs) note main' = main[OF this] then show "r = 0 \ degree r < degree g" by auto show "smult (coeff g (degree g) ^ (Suc (degree f) - degree g)) f = g * q + r" by (subst main'[THEN conjunct2, symmetric], simp add: degree_eq_length_coeffs, cases "f = 0"; cases "coeffs g", use g in auto) qed definition "pseudo_mod_main lc r d dr n = snd (pseudo_divmod_main lc 0 r d dr n)" lemma snd_pseudo_divmod_main: "snd (pseudo_divmod_main lc q r d dr n) = snd (pseudo_divmod_main lc q' r d dr n)" by (induct n arbitrary: q q' lc r d dr) (simp_all add: Let_def) definition pseudo_mod :: "'a::{comm_ring_1,semiring_1_no_zero_divisors} poly \ 'a poly \ 'a poly" where "pseudo_mod f g = snd (pseudo_divmod f g)" lemma pseudo_mod: fixes f g :: "'a::{comm_ring_1,semiring_1_no_zero_divisors} poly" defines "r \ pseudo_mod f g" assumes g: "g \ 0" shows "\a q. a \ 0 \ smult a f = g * q + r" "r = 0 \ degree r < degree g" proof - let ?cg = "coeff g (degree g)" let ?cge = "?cg ^ (Suc (degree f) - degree g)" define a where "a = ?cge" from r_def[unfolded pseudo_mod_def] obtain q where pdm: "pseudo_divmod f g = (q, r)" by (cases "pseudo_divmod f g") auto from pseudo_divmod[OF g pdm] have id: "smult a f = g * q + r" and "r = 0 \ degree r < degree g" by (auto simp: a_def) show "r = 0 \ degree r < degree g" by fact from g have "a \ 0" by (auto simp: a_def) with id show "\a q. a \ 0 \ smult a f = g * q + r" by auto qed lemma fst_pseudo_divmod_main_as_divide_poly_main: assumes d: "d \ 0" defines lc: "lc \ coeff d (degree d)" shows "fst (pseudo_divmod_main lc q r d dr n) = divide_poly_main lc (smult (lc^n) q) (smult (lc^n) r) d dr n" proof (induct n arbitrary: q r dr) case 0 then show ?case by simp next case (Suc n) note lc0 = leading_coeff_neq_0[OF d, folded lc] then have "pseudo_divmod_main lc q r d dr (Suc n) = pseudo_divmod_main lc (smult lc q + monom (coeff r dr) n) (smult lc r - monom (coeff r dr) n * d) d (dr - 1) n" by (simp add: Let_def ac_simps) also have "fst \ = divide_poly_main lc (smult (lc^n) (smult lc q + monom (coeff r dr) n)) (smult (lc^n) (smult lc r - monom (coeff r dr) n * d)) d (dr - 1) n" by (simp only: Suc[unfolded divide_poly_main.simps Let_def]) also have "\ = divide_poly_main lc (smult (lc ^ Suc n) q) (smult (lc ^ Suc n) r) d dr (Suc n)" unfolding smult_monom smult_distribs mult_smult_left[symmetric] using lc0 by (simp add: Let_def ac_simps) finally show ?case . qed subsubsection \Division in polynomials over fields\ lemma pseudo_divmod_field: fixes g :: "'a::field poly" assumes g: "g \ 0" and *: "pseudo_divmod f g = (q,r)" defines "c \ coeff g (degree g) ^ (Suc (degree f) - degree g)" shows "f = g * smult (1/c) q + smult (1/c) r" proof - from leading_coeff_neq_0[OF g] have c0: "c \ 0" by (auto simp: c_def) from pseudo_divmod(1)[OF g *, folded c_def] have "smult c f = g * q + r" by auto also have "smult (1 / c) \ = g * smult (1 / c) q + smult (1 / c) r" by (simp add: smult_add_right) finally show ?thesis using c0 by auto qed lemma divide_poly_main_field: fixes d :: "'a::field poly" assumes d: "d \ 0" defines lc: "lc \ coeff d (degree d)" shows "divide_poly_main lc q r d dr n = fst (pseudo_divmod_main lc (smult ((1 / lc)^n) q) (smult ((1 / lc)^n) r) d dr n)" unfolding lc by (subst fst_pseudo_divmod_main_as_divide_poly_main) (auto simp: d power_one_over) lemma divide_poly_field: fixes f g :: "'a::field poly" defines "f' \ smult ((1 / coeff g (degree g)) ^ (Suc (degree f) - degree g)) f" shows "f div g = fst (pseudo_divmod f' g)" proof (cases "g = 0") case True show ?thesis unfolding divide_poly_def pseudo_divmod_def Let_def f'_def True by (simp add: divide_poly_main_0) next case False from leading_coeff_neq_0[OF False] have "degree f' = degree f" by (auto simp: f'_def) then show ?thesis using length_coeffs_degree[of f'] length_coeffs_degree[of f] unfolding divide_poly_def pseudo_divmod_def Let_def divide_poly_main_field[OF False] length_coeffs_degree[OF False] f'_def by force qed instantiation poly :: ("{semidom_divide_unit_factor,idom_divide}") normalization_semidom begin definition unit_factor_poly :: "'a poly \ 'a poly" where "unit_factor_poly p = [:unit_factor (lead_coeff p):]" definition normalize_poly :: "'a poly \ 'a poly" where "normalize p = p div [:unit_factor (lead_coeff p):]" instance proof fix p :: "'a poly" show "unit_factor p * normalize p = p" proof (cases "p = 0") case True then show ?thesis by (simp add: unit_factor_poly_def normalize_poly_def) next case False then have "lead_coeff p \ 0" by simp then have *: "unit_factor (lead_coeff p) \ 0" using unit_factor_is_unit [of "lead_coeff p"] by auto then have "unit_factor (lead_coeff p) dvd 1" by (auto intro: unit_factor_is_unit) then have **: "unit_factor (lead_coeff p) dvd c" for c by (rule dvd_trans) simp have ***: "unit_factor (lead_coeff p) * (c div unit_factor (lead_coeff p)) = c" for c proof - from ** obtain b where "c = unit_factor (lead_coeff p) * b" .. with False * show ?thesis by simp qed have "p div [:unit_factor (lead_coeff p):] = map_poly (\c. c div unit_factor (lead_coeff p)) p" by (simp add: const_poly_dvd_iff div_const_poly_conv_map_poly **) then show ?thesis by (simp add: normalize_poly_def unit_factor_poly_def smult_conv_map_poly map_poly_map_poly o_def ***) qed next fix p :: "'a poly" assume "is_unit p" then obtain c where p: "p = [:c:]" "c dvd 1" by (auto simp: is_unit_poly_iff) then show "unit_factor p = p" by (simp add: unit_factor_poly_def monom_0 is_unit_unit_factor) next fix p :: "'a poly" assume "p \ 0" then show "is_unit (unit_factor p)" by (simp add: unit_factor_poly_def monom_0 is_unit_poly_iff unit_factor_is_unit) next fix a b :: "'a poly" assume "is_unit a" thus "unit_factor (a * b) = a * unit_factor b" by (auto simp: unit_factor_poly_def lead_coeff_mult unit_factor_mult elim!: is_unit_polyE) qed (simp_all add: normalize_poly_def unit_factor_poly_def monom_0 lead_coeff_mult unit_factor_mult) end instance poly :: ("{semidom_divide_unit_factor,idom_divide,normalization_semidom_multiplicative}") normalization_semidom_multiplicative by intro_classes (auto simp: unit_factor_poly_def lead_coeff_mult unit_factor_mult) lemma normalize_poly_eq_map_poly: "normalize p = map_poly (\x. x div unit_factor (lead_coeff p)) p" proof - have "[:unit_factor (lead_coeff p):] dvd p" by (metis unit_factor_poly_def unit_factor_self) then show ?thesis by (simp add: normalize_poly_def div_const_poly_conv_map_poly) qed lemma coeff_normalize [simp]: "coeff (normalize p) n = coeff p n div unit_factor (lead_coeff p)" by (simp add: normalize_poly_eq_map_poly coeff_map_poly) class field_unit_factor = field + unit_factor + assumes unit_factor_field [simp]: "unit_factor = id" begin subclass semidom_divide_unit_factor proof fix a assume "a \ 0" then have "1 = a * inverse a" by simp then have "a dvd 1" .. then show "unit_factor a dvd 1" by simp qed simp_all end lemma unit_factor_pCons: "unit_factor (pCons a p) = (if p = 0 then [:unit_factor a:] else unit_factor p)" by (simp add: unit_factor_poly_def) lemma normalize_monom [simp]: "normalize (monom a n) = monom (normalize a) n" by (cases "a = 0") (simp_all add: map_poly_monom normalize_poly_eq_map_poly degree_monom_eq) lemma unit_factor_monom [simp]: "unit_factor (monom a n) = [:unit_factor a:]" by (cases "a = 0") (simp_all add: unit_factor_poly_def degree_monom_eq) lemma normalize_const_poly: "normalize [:c:] = [:normalize c:]" by (simp add: normalize_poly_eq_map_poly map_poly_pCons) lemma normalize_smult: fixes c :: "'a :: {normalization_semidom_multiplicative, idom_divide}" shows "normalize (smult c p) = smult (normalize c) (normalize p)" proof - have "smult c p = [:c:] * p" by simp also have "normalize \ = smult (normalize c) (normalize p)" by (subst normalize_mult) (simp add: normalize_const_poly) finally show ?thesis . qed inductive eucl_rel_poly :: "'a::field poly \ 'a poly \ 'a poly \ 'a poly \ bool" where eucl_rel_poly_by0: "eucl_rel_poly x 0 (0, x)" | eucl_rel_poly_dividesI: "y \ 0 \ x = q * y \ eucl_rel_poly x y (q, 0)" | eucl_rel_poly_remainderI: "y \ 0 \ degree r < degree y \ x = q * y + r \ eucl_rel_poly x y (q, r)" lemma eucl_rel_poly_iff: "eucl_rel_poly x y (q, r) \ x = q * y + r \ (if y = 0 then q = 0 else r = 0 \ degree r < degree y)" by (auto elim: eucl_rel_poly.cases intro: eucl_rel_poly_by0 eucl_rel_poly_dividesI eucl_rel_poly_remainderI) lemma eucl_rel_poly_0: "eucl_rel_poly 0 y (0, 0)" by (simp add: eucl_rel_poly_iff) lemma eucl_rel_poly_by_0: "eucl_rel_poly x 0 (0, x)" by (simp add: eucl_rel_poly_iff) lemma eucl_rel_poly_pCons: assumes rel: "eucl_rel_poly x y (q, r)" assumes y: "y \ 0" assumes b: "b = coeff (pCons a r) (degree y) / coeff y (degree y)" shows "eucl_rel_poly (pCons a x) y (pCons b q, pCons a r - smult b y)" (is "eucl_rel_poly ?x y (?q, ?r)") proof - from assms have x: "x = q * y + r" and r: "r = 0 \ degree r < degree y" by (simp_all add: eucl_rel_poly_iff) from b x have "?x = ?q * y + ?r" by simp moreover have "?r = 0 \ degree ?r < degree y" proof (rule eq_zero_or_degree_less) show "degree ?r \ degree y" proof (rule degree_diff_le) from r show "degree (pCons a r) \ degree y" by auto show "degree (smult b y) \ degree y" by (rule degree_smult_le) qed from \y \ 0\ show "coeff ?r (degree y) = 0" by (simp add: b) qed ultimately show ?thesis unfolding eucl_rel_poly_iff using \y \ 0\ by simp qed lemma eucl_rel_poly_exists: "\q r. eucl_rel_poly x y (q, r)" proof (cases "y = 0") case False show ?thesis proof (induction x) case 0 then show ?case using eucl_rel_poly_0 by blast next case (pCons a x) then show ?case using False eucl_rel_poly_pCons by blast qed qed (use eucl_rel_poly_by0 in blast) lemma eucl_rel_poly_unique: assumes 1: "eucl_rel_poly x y (q1, r1)" assumes 2: "eucl_rel_poly x y (q2, r2)" shows "q1 = q2 \ r1 = r2" proof (cases "y = 0") assume "y = 0" with assms show ?thesis by (simp add: eucl_rel_poly_iff) next assume [simp]: "y \ 0" from 1 have q1: "x = q1 * y + r1" and r1: "r1 = 0 \ degree r1 < degree y" unfolding eucl_rel_poly_iff by simp_all from 2 have q2: "x = q2 * y + r2" and r2: "r2 = 0 \ degree r2 < degree y" unfolding eucl_rel_poly_iff by simp_all from q1 q2 have q3: "(q1 - q2) * y = r2 - r1" by (simp add: algebra_simps) from r1 r2 have r3: "(r2 - r1) = 0 \ degree (r2 - r1) < degree y" by (auto intro: degree_diff_less) show "q1 = q2 \ r1 = r2" proof (rule classical) assume "\ ?thesis" with q3 have "q1 \ q2" and "r1 \ r2" by auto with r3 have "degree (r2 - r1) < degree y" by simp also have "degree y \ degree (q1 - q2) + degree y" by simp also from \q1 \ q2\ have "\ = degree ((q1 - q2) * y)" by (simp add: degree_mult_eq) also from q3 have "\ = degree (r2 - r1)" by simp finally have "degree (r2 - r1) < degree (r2 - r1)" . then show ?thesis by simp qed qed lemma eucl_rel_poly_0_iff: "eucl_rel_poly 0 y (q, r) \ q = 0 \ r = 0" by (auto dest: eucl_rel_poly_unique intro: eucl_rel_poly_0) lemma eucl_rel_poly_by_0_iff: "eucl_rel_poly x 0 (q, r) \ q = 0 \ r = x" by (auto dest: eucl_rel_poly_unique intro: eucl_rel_poly_by_0) lemmas eucl_rel_poly_unique_div = eucl_rel_poly_unique [THEN conjunct1] lemmas eucl_rel_poly_unique_mod = eucl_rel_poly_unique [THEN conjunct2] instantiation poly :: (field) semidom_modulo begin definition modulo_poly :: "'a poly \ 'a poly \ 'a poly" where mod_poly_def: "f mod g = (if g = 0 then f else pseudo_mod (smult ((1 / lead_coeff g) ^ (Suc (degree f) - degree g)) f) g)" instance proof fix x y :: "'a poly" show "x div y * y + x mod y = x" proof (cases "y = 0") case True then show ?thesis by (simp add: divide_poly_0 mod_poly_def) next case False then have "pseudo_divmod (smult ((1 / lead_coeff y) ^ (Suc (degree x) - degree y)) x) y = (x div y, x mod y)" by (simp add: divide_poly_field mod_poly_def pseudo_mod_def) with False pseudo_divmod [OF False this] show ?thesis by (simp add: power_mult_distrib [symmetric] ac_simps) qed qed end lemma eucl_rel_poly: "eucl_rel_poly x y (x div y, x mod y)" unfolding eucl_rel_poly_iff proof show "x = x div y * y + x mod y" by (simp add: div_mult_mod_eq) show "if y = 0 then x div y = 0 else x mod y = 0 \ degree (x mod y) < degree y" proof (cases "y = 0") case True then show ?thesis by auto next case False with pseudo_mod[OF this] show ?thesis by (simp add: mod_poly_def) qed qed lemma div_poly_eq: "eucl_rel_poly x y (q, r) \ x div y = q" for x :: "'a::field poly" by (rule eucl_rel_poly_unique_div [OF eucl_rel_poly]) lemma mod_poly_eq: "eucl_rel_poly x y (q, r) \ x mod y = r" for x :: "'a::field poly" by (rule eucl_rel_poly_unique_mod [OF eucl_rel_poly]) instance poly :: (field) idom_modulo .. lemma div_pCons_eq: "pCons a p div q = (if q = 0 then 0 else pCons (coeff (pCons a (p mod q)) (degree q) / lead_coeff q) (p div q))" using eucl_rel_poly_pCons [OF eucl_rel_poly _ refl, of q a p] by (auto intro: div_poly_eq) lemma mod_pCons_eq: "pCons a p mod q = (if q = 0 then pCons a p else pCons a (p mod q) - smult (coeff (pCons a (p mod q)) (degree q) / lead_coeff q) q)" using eucl_rel_poly_pCons [OF eucl_rel_poly _ refl, of q a p] by (auto intro: mod_poly_eq) lemma div_mod_fold_coeffs: "(p div q, p mod q) = (if q = 0 then (0, p) else fold_coeffs (\a (s, r). let b = coeff (pCons a r) (degree q) / coeff q (degree q) in (pCons b s, pCons a r - smult b q)) p (0, 0))" by (rule sym, induct p) (auto simp: div_pCons_eq mod_pCons_eq Let_def) lemma degree_mod_less: "y \ 0 \ x mod y = 0 \ degree (x mod y) < degree y" using eucl_rel_poly [of x y] unfolding eucl_rel_poly_iff by simp lemma degree_mod_less': "b \ 0 \ a mod b \ 0 \ degree (a mod b) < degree b" using degree_mod_less[of b a] by auto lemma div_poly_less: fixes x :: "'a::field poly" assumes "degree x < degree y" shows "x div y = 0" proof - from assms have "eucl_rel_poly x y (0, x)" by (simp add: eucl_rel_poly_iff) then show "x div y = 0" by (rule div_poly_eq) qed lemma mod_poly_less: assumes "degree x < degree y" shows "x mod y = x" proof - from assms have "eucl_rel_poly x y (0, x)" by (simp add: eucl_rel_poly_iff) then show "x mod y = x" by (rule mod_poly_eq) qed lemma eucl_rel_poly_smult_left: "eucl_rel_poly x y (q, r) \ eucl_rel_poly (smult a x) y (smult a q, smult a r)" by (simp add: eucl_rel_poly_iff smult_add_right) lemma div_smult_left: "(smult a x) div y = smult a (x div y)" for x y :: "'a::field poly" by (rule div_poly_eq, rule eucl_rel_poly_smult_left, rule eucl_rel_poly) lemma mod_smult_left: "(smult a x) mod y = smult a (x mod y)" by (rule mod_poly_eq, rule eucl_rel_poly_smult_left, rule eucl_rel_poly) lemma poly_div_minus_left [simp]: "(- x) div y = - (x div y)" for x y :: "'a::field poly" using div_smult_left [of "- 1::'a"] by simp lemma poly_mod_minus_left [simp]: "(- x) mod y = - (x mod y)" for x y :: "'a::field poly" using mod_smult_left [of "- 1::'a"] by simp lemma eucl_rel_poly_add_left: assumes "eucl_rel_poly x y (q, r)" assumes "eucl_rel_poly x' y (q', r')" shows "eucl_rel_poly (x + x') y (q + q', r + r')" using assms unfolding eucl_rel_poly_iff by (auto simp: algebra_simps degree_add_less) lemma poly_div_add_left: "(x + y) div z = x div z + y div z" for x y z :: "'a::field poly" using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly] by (rule div_poly_eq) lemma poly_mod_add_left: "(x + y) mod z = x mod z + y mod z" for x y z :: "'a::field poly" using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly] by (rule mod_poly_eq) lemma poly_div_diff_left: "(x - y) div z = x div z - y div z" for x y z :: "'a::field poly" by (simp only: diff_conv_add_uminus poly_div_add_left poly_div_minus_left) lemma poly_mod_diff_left: "(x - y) mod z = x mod z - y mod z" for x y z :: "'a::field poly" by (simp only: diff_conv_add_uminus poly_mod_add_left poly_mod_minus_left) lemma eucl_rel_poly_smult_right: "a \ 0 \ eucl_rel_poly x y (q, r) \ eucl_rel_poly x (smult a y) (smult (inverse a) q, r)" by (simp add: eucl_rel_poly_iff) lemma div_smult_right: "a \ 0 \ x div (smult a y) = smult (inverse a) (x div y)" for x y :: "'a::field poly" by (rule div_poly_eq, erule eucl_rel_poly_smult_right, rule eucl_rel_poly) lemma mod_smult_right: "a \ 0 \ x mod (smult a y) = x mod y" by (rule mod_poly_eq, erule eucl_rel_poly_smult_right, rule eucl_rel_poly) lemma poly_div_minus_right [simp]: "x div (- y) = - (x div y)" for x y :: "'a::field poly" using div_smult_right [of "- 1::'a"] by (simp add: nonzero_inverse_minus_eq) lemma poly_mod_minus_right [simp]: "x mod (- y) = x mod y" for x y :: "'a::field poly" using mod_smult_right [of "- 1::'a"] by simp lemma eucl_rel_poly_mult: assumes "eucl_rel_poly x y (q, r)" "eucl_rel_poly q z (q', r')" shows "eucl_rel_poly x (y * z) (q', y * r' + r)" proof (cases "y = 0") case True with assms eucl_rel_poly_0_iff show ?thesis by (force simp add: eucl_rel_poly_iff) next case False show ?thesis proof (cases "r' = 0") case True with assms show ?thesis by (auto simp add: eucl_rel_poly_iff degree_mult_eq) next case False with assms \y \ 0\ show ?thesis by (auto simp add: eucl_rel_poly_iff degree_add_less degree_mult_eq field_simps) qed qed lemma poly_div_mult_right: "x div (y * z) = (x div y) div z" for x y z :: "'a::field poly" by (rule div_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+) lemma poly_mod_mult_right: "x mod (y * z) = y * (x div y mod z) + x mod y" for x y z :: "'a::field poly" by (rule mod_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+) lemma mod_pCons: fixes a :: "'a::field" and x y :: "'a::field poly" assumes y: "y \ 0" defines "b \ coeff (pCons a (x mod y)) (degree y) / coeff y (degree y)" shows "(pCons a x) mod y = pCons a (x mod y) - smult b y" unfolding b_def by (rule mod_poly_eq, rule eucl_rel_poly_pCons [OF eucl_rel_poly y refl]) subsubsection \List-based versions for fast implementation\ (* Subsection by: Sebastiaan Joosten René Thiemann Akihisa Yamada *) fun minus_poly_rev_list :: "'a :: group_add list \ 'a list \ 'a list" where "minus_poly_rev_list (x # xs) (y # ys) = (x - y) # (minus_poly_rev_list xs ys)" | "minus_poly_rev_list xs [] = xs" | "minus_poly_rev_list [] (y # ys) = []" fun pseudo_divmod_main_list :: "'a::comm_ring_1 \ 'a list \ 'a list \ 'a list \ nat \ 'a list \ 'a list" where "pseudo_divmod_main_list lc q r d (Suc n) = (let rr = map ((*) lc) r; a = hd r; qqq = cCons a (map ((*) lc) q); rrr = tl (if a = 0 then rr else minus_poly_rev_list rr (map ((*) a) d)) in pseudo_divmod_main_list lc qqq rrr d n)" | "pseudo_divmod_main_list lc q r d 0 = (q, r)" fun pseudo_mod_main_list :: "'a::comm_ring_1 \ 'a list \ 'a list \ nat \ 'a list" where "pseudo_mod_main_list lc r d (Suc n) = (let rr = map ((*) lc) r; a = hd r; rrr = tl (if a = 0 then rr else minus_poly_rev_list rr (map ((*) a) d)) in pseudo_mod_main_list lc rrr d n)" | "pseudo_mod_main_list lc r d 0 = r" fun divmod_poly_one_main_list :: "'a::comm_ring_1 list \ 'a list \ 'a list \ nat \ 'a list \ 'a list" where "divmod_poly_one_main_list q r d (Suc n) = (let a = hd r; qqq = cCons a q; rr = tl (if a = 0 then r else minus_poly_rev_list r (map ((*) a) d)) in divmod_poly_one_main_list qqq rr d n)" | "divmod_poly_one_main_list q r d 0 = (q, r)" fun mod_poly_one_main_list :: "'a::comm_ring_1 list \ 'a list \ nat \ 'a list" where "mod_poly_one_main_list r d (Suc n) = (let a = hd r; rr = tl (if a = 0 then r else minus_poly_rev_list r (map ((*) a) d)) in mod_poly_one_main_list rr d n)" | "mod_poly_one_main_list r d 0 = r" definition pseudo_divmod_list :: "'a::comm_ring_1 list \ 'a list \ 'a list \ 'a list" where "pseudo_divmod_list p q = (if q = [] then ([], p) else (let rq = rev q; (qu,re) = pseudo_divmod_main_list (hd rq) [] (rev p) rq (1 + length p - length q) in (qu, rev re)))" definition pseudo_mod_list :: "'a::comm_ring_1 list \ 'a list \ 'a list" where "pseudo_mod_list p q = (if q = [] then p else (let rq = rev q; re = pseudo_mod_main_list (hd rq) (rev p) rq (1 + length p - length q) in rev re))" lemma minus_zero_does_nothing: "minus_poly_rev_list x (map ((*) 0) y) = x" for x :: "'a::ring list" by (induct x y rule: minus_poly_rev_list.induct) auto lemma length_minus_poly_rev_list [simp]: "length (minus_poly_rev_list xs ys) = length xs" by (induct xs ys rule: minus_poly_rev_list.induct) auto lemma if_0_minus_poly_rev_list: "(if a = 0 then x else minus_poly_rev_list x (map ((*) a) y)) = minus_poly_rev_list x (map ((*) a) y)" for a :: "'a::ring" by(cases "a = 0") (simp_all add: minus_zero_does_nothing) lemma Poly_append: "Poly (a @ b) = Poly a + monom 1 (length a) * Poly b" for a :: "'a::comm_semiring_1 list" by (induct a) (auto simp: monom_0 monom_Suc) lemma minus_poly_rev_list: "length p \ length q \ Poly (rev (minus_poly_rev_list (rev p) (rev q))) = Poly p - monom 1 (length p - length q) * Poly q" for p q :: "'a :: comm_ring_1 list" proof (induct "rev p" "rev q" arbitrary: p q rule: minus_poly_rev_list.induct) case (1 x xs y ys) then have "length (rev q) \ length (rev p)" by simp from this[folded 1(2,3)] have ys_xs: "length ys \ length xs" by simp then have *: "Poly (rev (minus_poly_rev_list xs ys)) = Poly (rev xs) - monom 1 (length xs - length ys) * Poly (rev ys)" by (subst "1.hyps"(1)[of "rev xs" "rev ys", unfolded rev_rev_ident length_rev]) auto have "Poly p - monom 1 (length p - length q) * Poly q = Poly (rev (rev p)) - monom 1 (length (rev (rev p)) - length (rev (rev q))) * Poly (rev (rev q))" by simp also have "\ = Poly (rev (x # xs)) - monom 1 (length (x # xs) - length (y # ys)) * Poly (rev (y # ys))" unfolding 1(2,3) by simp also from ys_xs have "\ = Poly (rev xs) + monom x (length xs) - (monom 1 (length xs - length ys) * Poly (rev ys) + monom y (length xs))" by (simp add: Poly_append distrib_left mult_monom smult_monom) also have "\ = Poly (rev (minus_poly_rev_list xs ys)) + monom (x - y) (length xs)" unfolding * diff_monom[symmetric] by simp finally show ?case by (simp add: 1(2,3)[symmetric] smult_monom Poly_append) qed auto lemma smult_monom_mult: "smult a (monom b n * f) = monom (a * b) n * f" using smult_monom [of a _ n] by (metis mult_smult_left) lemma head_minus_poly_rev_list: "length d \ length r \ d \ [] \ hd (minus_poly_rev_list (map ((*) (last d)) r) (map ((*) (hd r)) (rev d))) = 0" for d r :: "'a::comm_ring list" proof (induct r) case Nil then show ?case by simp next case (Cons a rs) then show ?case by (cases "rev d") (simp_all add: ac_simps) qed lemma Poly_map: "Poly (map ((*) a) p) = smult a (Poly p)" proof (induct p) case Nil then show ?case by simp next case (Cons x xs) then show ?case by (cases "Poly xs = 0") auto qed lemma last_coeff_is_hd: "xs \ [] \ coeff (Poly xs) (length xs - 1) = hd (rev xs)" by (simp_all add: hd_conv_nth rev_nth nth_default_nth nth_append) lemma pseudo_divmod_main_list_invar: assumes leading_nonzero: "last d \ 0" and lc: "last d = lc" and "d \ []" and "pseudo_divmod_main_list lc q (rev r) (rev d) n = (q', rev r')" and "n = 1 + length r - length d" shows "pseudo_divmod_main lc (monom 1 n * Poly q) (Poly r) (Poly d) (length r - 1) n = (Poly q', Poly r')" using assms(4-) proof (induct n arbitrary: r q) case (Suc n) from Suc.prems have *: "\ Suc (length r) \ length d" by simp with \d \ []\ have "r \ []" using Suc_leI length_greater_0_conv list.size(3) by fastforce let ?a = "(hd (rev r))" let ?rr = "map ((*) lc) (rev r)" let ?rrr = "rev (tl (minus_poly_rev_list ?rr (map ((*) ?a) (rev d))))" let ?qq = "cCons ?a (map ((*) lc) q)" from * Suc(3) have n: "n = (1 + length r - length d - 1)" by simp from * have rr_val:"(length ?rrr) = (length r - 1)" by auto with \r \ []\ * have rr_smaller: "(1 + length r - length d - 1) = (1 + length ?rrr - length d)" by auto from * have id: "Suc (length r) - length d = Suc (length r - length d)" by auto from Suc.prems * have "pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) (1 + length r - length d - 1) = (q', rev r')" by (simp add: Let_def if_0_minus_poly_rev_list id) with n have v: "pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) n = (q', rev r')" by auto from * have sucrr:"Suc (length r) - length d = Suc (length r - length d)" using Suc_diff_le not_less_eq_eq by blast from Suc(3) \r \ []\ have n_ok : "n = 1 + (length ?rrr) - length d" by simp have cong: "\x1 x2 x3 x4 y1 y2 y3 y4. x1 = y1 \ x2 = y2 \ x3 = y3 \ x4 = y4 \ pseudo_divmod_main lc x1 x2 x3 x4 n = pseudo_divmod_main lc y1 y2 y3 y4 n" by simp have hd_rev: "coeff (Poly r) (length r - Suc 0) = hd (rev r)" using last_coeff_is_hd[OF \r \ []\] by simp show ?case unfolding Suc.hyps(1)[OF v n_ok, symmetric] pseudo_divmod_main.simps Let_def proof (rule cong[OF _ _ refl], goal_cases) case 1 show ?case by (simp add: monom_Suc hd_rev[symmetric] smult_monom Poly_map) next case 2 show ?case proof (subst Poly_on_rev_starting_with_0, goal_cases) show "hd (minus_poly_rev_list (map ((*) lc) (rev r)) (map ((*) (hd (rev r))) (rev d))) = 0" by (fold lc, subst head_minus_poly_rev_list, insert * \d \ []\, auto) from * have "length d \ length r" by simp then show "smult lc (Poly r) - monom (coeff (Poly r) (length r - 1)) n * Poly d = Poly (rev (minus_poly_rev_list (map ((*) lc) (rev r)) (map ((*) (hd (rev r))) (rev d))))" by (fold rev_map) (auto simp add: n smult_monom_mult Poly_map hd_rev [symmetric] minus_poly_rev_list) qed qed simp qed simp lemma pseudo_divmod_impl [code]: "pseudo_divmod f g = map_prod poly_of_list poly_of_list (pseudo_divmod_list (coeffs f) (coeffs g))" for f g :: "'a::comm_ring_1 poly" proof (cases "g = 0") case False then have "last (coeffs g) \ 0" and "last (coeffs g) = lead_coeff g" and "coeffs g \ []" by (simp_all add: last_coeffs_eq_coeff_degree) moreover obtain q r where qr: "pseudo_divmod_main_list (last (coeffs g)) (rev []) (rev (coeffs f)) (rev (coeffs g)) (1 + length (coeffs f) - length (coeffs g)) = (q, rev (rev r))" by force ultimately have "(Poly q, Poly (rev r)) = pseudo_divmod_main (lead_coeff g) 0 f g (length (coeffs f) - Suc 0) (Suc (length (coeffs f)) - length (coeffs g))" by (subst pseudo_divmod_main_list_invar [symmetric]) auto moreover have "pseudo_divmod_main_list (hd (rev (coeffs g))) [] (rev (coeffs f)) (rev (coeffs g)) (1 + length (coeffs f) - length (coeffs g)) = (q, r)" using qr hd_rev [OF \coeffs g \ []\] by simp ultimately show ?thesis by (auto simp: degree_eq_length_coeffs pseudo_divmod_def pseudo_divmod_list_def Let_def) next case True then show ?thesis by (auto simp add: pseudo_divmod_def pseudo_divmod_list_def) qed lemma pseudo_mod_main_list: "snd (pseudo_divmod_main_list l q xs ys n) = pseudo_mod_main_list l xs ys n" by (induct n arbitrary: l q xs ys) (auto simp: Let_def) lemma pseudo_mod_impl[code]: "pseudo_mod f g = poly_of_list (pseudo_mod_list (coeffs f) (coeffs g))" proof - have snd_case: "\f g p. snd ((\(x,y). (f x, g y)) p) = g (snd p)" by auto show ?thesis unfolding pseudo_mod_def pseudo_divmod_impl pseudo_divmod_list_def pseudo_mod_list_def Let_def by (simp add: snd_case pseudo_mod_main_list) qed subsubsection \Improved Code-Equations for Polynomial (Pseudo) Division\ lemma pdivmod_pdivmodrel: "eucl_rel_poly p q (r, s) \ (p div q, p mod q) = (r, s)" by (metis eucl_rel_poly eucl_rel_poly_unique) lemma pdivmod_via_pseudo_divmod: "(f div g, f mod g) = (if g = 0 then (0, f) else let ilc = inverse (coeff g (degree g)); h = smult ilc g; (q,r) = pseudo_divmod f h in (smult ilc q, r))" (is "?l = ?r") proof (cases "g = 0") case True then show ?thesis by simp next case False define lc where "lc = inverse (coeff g (degree g))" define h where "h = smult lc g" from False have h1: "coeff h (degree h) = 1" and lc: "lc \ 0" by (auto simp: h_def lc_def) then have h0: "h \ 0" by auto obtain q r where p: "pseudo_divmod f h = (q, r)" by force from False have id: "?r = (smult lc q, r)" by (auto simp: Let_def h_def[symmetric] lc_def[symmetric] p) from pseudo_divmod[OF h0 p, unfolded h1] have f: "f = h * q + r" and r: "r = 0 \ degree r < degree h" by auto from f r h0 have "eucl_rel_poly f h (q, r)" by (auto simp: eucl_rel_poly_iff) then have "(f div h, f mod h) = (q, r)" by (simp add: pdivmod_pdivmodrel) with lc have "(f div g, f mod g) = (smult lc q, r)" by (auto simp: h_def div_smult_right[OF lc] mod_smult_right[OF lc]) with id show ?thesis by auto qed lemma pdivmod_via_pseudo_divmod_list: "(f div g, f mod g) = (let cg = coeffs g in if cg = [] then (0, f) else let cf = coeffs f; ilc = inverse (last cg); ch = map ((*) ilc) cg; (q, r) = pseudo_divmod_main_list 1 [] (rev cf) (rev ch) (1 + length cf - length cg) in (poly_of_list (map ((*) ilc) q), poly_of_list (rev r)))" proof - note d = pdivmod_via_pseudo_divmod pseudo_divmod_impl pseudo_divmod_list_def show ?thesis proof (cases "g = 0") case True with d show ?thesis by auto next case False define ilc where "ilc = inverse (coeff g (degree g))" from False have ilc: "ilc \ 0" by (auto simp: ilc_def) with False have id: "g = 0 \ False" "coeffs g = [] \ False" "last (coeffs g) = coeff g (degree g)" "coeffs (smult ilc g) = [] \ False" by (auto simp: last_coeffs_eq_coeff_degree) have id2: "hd (rev (coeffs (smult ilc g))) = 1" by (subst hd_rev, insert id ilc, auto simp: coeffs_smult, subst last_map, auto simp: id ilc_def) have id3: "length (coeffs (smult ilc g)) = length (coeffs g)" "rev (coeffs (smult ilc g)) = rev (map ((*) ilc) (coeffs g))" unfolding coeffs_smult using ilc by auto obtain q r where pair: "pseudo_divmod_main_list 1 [] (rev (coeffs f)) (rev (map ((*) ilc) (coeffs g))) (1 + length (coeffs f) - length (coeffs g)) = (q, r)" by force show ?thesis unfolding d Let_def id if_False ilc_def[symmetric] map_prod_def[symmetric] id2 unfolding id3 pair map_prod_def split by (auto simp: Poly_map) qed qed lemma pseudo_divmod_main_list_1: "pseudo_divmod_main_list 1 = divmod_poly_one_main_list" proof (intro ext, goal_cases) case (1 q r d n) have *: "map ((*) 1) xs = xs" for xs :: "'a list" by (induct xs) auto show ?case by (induct n arbitrary: q r d) (auto simp: * Let_def) qed fun divide_poly_main_list :: "'a::idom_divide \ 'a list \ 'a list \ 'a list \ nat \ 'a list" where "divide_poly_main_list lc q r d (Suc n) = (let cr = hd r in if cr = 0 then divide_poly_main_list lc (cCons cr q) (tl r) d n else let a = cr div lc; qq = cCons a q; rr = minus_poly_rev_list r (map ((*) a) d) in if hd rr = 0 then divide_poly_main_list lc qq (tl rr) d n else [])" | "divide_poly_main_list lc q r d 0 = q" lemma divide_poly_main_list_simp [simp]: "divide_poly_main_list lc q r d (Suc n) = (let cr = hd r; a = cr div lc; qq = cCons a q; rr = minus_poly_rev_list r (map ((*) a) d) in if hd rr = 0 then divide_poly_main_list lc qq (tl rr) d n else [])" by (simp add: Let_def minus_zero_does_nothing) declare divide_poly_main_list.simps(1)[simp del] definition divide_poly_list :: "'a::idom_divide poly \ 'a poly \ 'a poly" where "divide_poly_list f g = (let cg = coeffs g in if cg = [] then g else let cf = coeffs f; cgr = rev cg in poly_of_list (divide_poly_main_list (hd cgr) [] (rev cf) cgr (1 + length cf - length cg)))" lemmas pdivmod_via_divmod_list = pdivmod_via_pseudo_divmod_list[unfolded pseudo_divmod_main_list_1] lemma mod_poly_one_main_list: "snd (divmod_poly_one_main_list q r d n) = mod_poly_one_main_list r d n" by (induct n arbitrary: q r d) (auto simp: Let_def) lemma mod_poly_code [code]: "f mod g = (let cg = coeffs g in if cg = [] then f else let cf = coeffs f; ilc = inverse (last cg); ch = map ((*) ilc) cg; r = mod_poly_one_main_list (rev cf) (rev ch) (1 + length cf - length cg) in poly_of_list (rev r))" (is "_ = ?rhs") proof - have "snd (f div g, f mod g) = ?rhs" unfolding pdivmod_via_divmod_list Let_def mod_poly_one_main_list [symmetric, of _ _ _ Nil] by (auto split: prod.splits) then show ?thesis by simp qed definition div_field_poly_impl :: "'a :: field poly \ 'a poly \ 'a poly" where "div_field_poly_impl f g = (let cg = coeffs g in if cg = [] then 0 else let cf = coeffs f; ilc = inverse (last cg); ch = map ((*) ilc) cg; q = fst (divmod_poly_one_main_list [] (rev cf) (rev ch) (1 + length cf - length cg)) in poly_of_list ((map ((*) ilc) q)))" text \We do not declare the following lemma as code equation, since then polynomial division on non-fields will no longer be executable. However, a code-unfold is possible, since \div_field_poly_impl\ is a bit more efficient than the generic polynomial division.\ lemma div_field_poly_impl[code_unfold]: "(div) = div_field_poly_impl" proof (intro ext) fix f g :: "'a poly" have "fst (f div g, f mod g) = div_field_poly_impl f g" unfolding div_field_poly_impl_def pdivmod_via_divmod_list Let_def by (auto split: prod.splits) then show "f div g = div_field_poly_impl f g" by simp qed lemma divide_poly_main_list: assumes lc0: "lc \ 0" and lc: "last d = lc" and d: "d \ []" and "n = (1 + length r - length d)" shows "Poly (divide_poly_main_list lc q (rev r) (rev d) n) = divide_poly_main lc (monom 1 n * Poly q) (Poly r) (Poly d) (length r - 1) n" using assms(4-) proof (induct "n" arbitrary: r q) case (Suc n) from Suc.prems have ifCond: "\ Suc (length r) \ length d" by simp with d have r: "r \ []" using Suc_leI length_greater_0_conv list.size(3) by fastforce then obtain rr lcr where r: "r = rr @ [lcr]" by (cases r rule: rev_cases) auto from d lc obtain dd where d: "d = dd @ [lc]" by (cases d rule: rev_cases) auto from Suc(2) ifCond have n: "n = 1 + length rr - length d" by (auto simp: r) from ifCond have len: "length dd \ length rr" by (simp add: r d) show ?case proof (cases "lcr div lc * lc = lcr") case False with r d show ?thesis unfolding Suc(2)[symmetric] by (auto simp add: Let_def nth_default_append) next case True with r d have id: "?thesis \ Poly (divide_poly_main_list lc (cCons (lcr div lc) q) (rev (rev (minus_poly_rev_list (rev rr) (rev (map ((*) (lcr div lc)) dd))))) (rev d) n) = divide_poly_main lc (monom 1 (Suc n) * Poly q + monom (lcr div lc) n) (Poly r - monom (lcr div lc) n * Poly d) (Poly d) (length rr - 1) n" by (cases r rule: rev_cases; cases "d" rule: rev_cases) (auto simp add: Let_def rev_map nth_default_append) have cong: "\x1 x2 x3 x4 y1 y2 y3 y4. x1 = y1 \ x2 = y2 \ x3 = y3 \ x4 = y4 \ divide_poly_main lc x1 x2 x3 x4 n = divide_poly_main lc y1 y2 y3 y4 n" by simp show ?thesis unfolding id proof (subst Suc(1), simp add: n, subst minus_poly_rev_list, force simp: len, rule cong[OF _ _ refl], goal_cases) case 2 have "monom lcr (length rr) = monom (lcr div lc) (length rr - length dd) * monom lc (length dd)" by (simp add: mult_monom len True) then show ?case unfolding r d Poly_append n ring_distribs by (auto simp: Poly_map smult_monom smult_monom_mult) qed (auto simp: len monom_Suc smult_monom) qed qed simp lemma divide_poly_list[code]: "f div g = divide_poly_list f g" proof - note d = divide_poly_def divide_poly_list_def show ?thesis proof (cases "g = 0") case True show ?thesis by (auto simp: d True) next case False then obtain cg lcg where cg: "coeffs g = cg @ [lcg]" by (cases "coeffs g" rule: rev_cases) auto with False have id: "(g = 0) = False" "(cg @ [lcg] = []) = False" by auto from cg False have lcg: "coeff g (degree g) = lcg" using last_coeffs_eq_coeff_degree last_snoc by force with False have "lcg \ 0" by auto from cg Poly_coeffs [of g] have ltp: "Poly (cg @ [lcg]) = g" by auto show ?thesis unfolding d cg Let_def id if_False poly_of_list_def by (subst divide_poly_main_list, insert False cg \lcg \ 0\) (auto simp: lcg ltp, simp add: degree_eq_length_coeffs) qed qed subsection \Primality and irreducibility in polynomial rings\ lemma prod_mset_const_poly: "(\x\#A. [:f x:]) = [:prod_mset (image_mset f A):]" by (induct A) (simp_all add: ac_simps) lemma irreducible_const_poly_iff: fixes c :: "'a :: {comm_semiring_1,semiring_no_zero_divisors}" shows "irreducible [:c:] \ irreducible c" proof assume A: "irreducible c" show "irreducible [:c:]" proof (rule irreducibleI) fix a b assume ab: "[:c:] = a * b" hence "degree [:c:] = degree (a * b)" by (simp only: ) also from A ab have "a \ 0" "b \ 0" by auto hence "degree (a * b) = degree a + degree b" by (simp add: degree_mult_eq) finally have "degree a = 0" "degree b = 0" by auto then obtain a' b' where ab': "a = [:a':]" "b = [:b':]" by (auto elim!: degree_eq_zeroE) from ab have "coeff [:c:] 0 = coeff (a * b) 0" by (simp only: ) hence "c = a' * b'" by (simp add: ab' mult_ac) from A and this have "a' dvd 1 \ b' dvd 1" by (rule irreducibleD) with ab' show "a dvd 1 \ b dvd 1" by (auto simp add: is_unit_const_poly_iff) qed (insert A, auto simp: irreducible_def is_unit_poly_iff) next assume A: "irreducible [:c:]" then have "c \ 0" and "\ c dvd 1" by (auto simp add: irreducible_def is_unit_const_poly_iff) then show "irreducible c" proof (rule irreducibleI) fix a b assume ab: "c = a * b" hence "[:c:] = [:a:] * [:b:]" by (simp add: mult_ac) from A and this have "[:a:] dvd 1 \ [:b:] dvd 1" by (rule irreducibleD) then show "a dvd 1 \ b dvd 1" by (auto simp add: is_unit_const_poly_iff) qed qed lemma lift_prime_elem_poly: assumes "prime_elem (c :: 'a :: semidom)" shows "prime_elem [:c:]" proof (rule prime_elemI) fix a b assume *: "[:c:] dvd a * b" from * have dvd: "c dvd coeff (a * b) n" for n by (subst (asm) const_poly_dvd_iff) blast { define m where "m = (GREATEST m. \c dvd coeff b m)" assume "\[:c:] dvd b" hence A: "\i. \c dvd coeff b i" by (subst (asm) const_poly_dvd_iff) blast have B: "\i. \c dvd coeff b i \ i \ degree b" by (auto intro: le_degree) have coeff_m: "\c dvd coeff b m" unfolding m_def by (rule GreatestI_ex_nat[OF A B]) have "i \ m" if "\c dvd coeff b i" for i unfolding m_def by (metis (mono_tags, lifting) B Greatest_le_nat that) hence dvd_b: "c dvd coeff b i" if "i > m" for i using that by force have "c dvd coeff a i" for i proof (induction i rule: nat_descend_induct[of "degree a"]) case (base i) thus ?case by (simp add: coeff_eq_0) next case (descend i) let ?A = "{..i+m} - {i}" have "c dvd coeff (a * b) (i + m)" by (rule dvd) also have "coeff (a * b) (i + m) = (\k\i + m. coeff a k * coeff b (i + m - k))" by (simp add: coeff_mult) also have "{..i+m} = insert i ?A" by auto also have "(\k\\. coeff a k * coeff b (i + m - k)) = coeff a i * coeff b m + (\k\?A. coeff a k * coeff b (i + m - k))" (is "_ = _ + ?S") by (subst sum.insert) simp_all finally have eq: "c dvd coeff a i * coeff b m + ?S" . moreover have "c dvd ?S" proof (rule dvd_sum) fix k assume k: "k \ {..i+m} - {i}" show "c dvd coeff a k * coeff b (i + m - k)" proof (cases "k < i") case False with k have "c dvd coeff a k" by (intro descend.IH) simp thus ?thesis by simp next case True hence "c dvd coeff b (i + m - k)" by (intro dvd_b) simp thus ?thesis by simp qed qed ultimately have "c dvd coeff a i * coeff b m" by (simp add: dvd_add_left_iff) with assms coeff_m show "c dvd coeff a i" by (simp add: prime_elem_dvd_mult_iff) qed hence "[:c:] dvd a" by (subst const_poly_dvd_iff) blast } then show "[:c:] dvd a \ [:c:] dvd b" by blast next from assms show "[:c:] \ 0" and "\ [:c:] dvd 1" by (simp_all add: prime_elem_def is_unit_const_poly_iff) qed lemma prime_elem_const_poly_iff: fixes c :: "'a :: semidom" shows "prime_elem [:c:] \ prime_elem c" proof assume A: "prime_elem [:c:]" show "prime_elem c" proof (rule prime_elemI) fix a b assume "c dvd a * b" hence "[:c:] dvd [:a:] * [:b:]" by (simp add: mult_ac) from A and this have "[:c:] dvd [:a:] \ [:c:] dvd [:b:]" by (rule prime_elem_dvd_multD) thus "c dvd a \ c dvd b" by simp qed (insert A, auto simp: prime_elem_def is_unit_poly_iff) qed (auto intro: lift_prime_elem_poly) subsection \Content and primitive part of a polynomial\ definition content :: "'a::semiring_gcd poly \ 'a" where "content p = gcd_list (coeffs p)" lemma content_eq_fold_coeffs [code]: "content p = fold_coeffs gcd p 0" by (simp add: content_def Gcd_fin.set_eq_fold fold_coeffs_def foldr_fold fun_eq_iff ac_simps) lemma content_0 [simp]: "content 0 = 0" by (simp add: content_def) lemma content_1 [simp]: "content 1 = 1" by (simp add: content_def) lemma content_const [simp]: "content [:c:] = normalize c" by (simp add: content_def cCons_def) lemma const_poly_dvd_iff_dvd_content: "[:c:] dvd p \ c dvd content p" for c :: "'a::semiring_gcd" proof (cases "p = 0") case True then show ?thesis by simp next case False have "[:c:] dvd p \ (\n. c dvd coeff p n)" by (rule const_poly_dvd_iff) also have "\ \ (\a\set (coeffs p). c dvd a)" proof safe fix n :: nat assume "\a\set (coeffs p). c dvd a" then show "c dvd coeff p n" by (cases "n \ degree p") (auto simp: coeff_eq_0 coeffs_def split: if_splits) qed (auto simp: coeffs_def simp del: upt_Suc split: if_splits) also have "\ \ c dvd content p" by (simp add: content_def dvd_Gcd_fin_iff dvd_mult_unit_iff) finally show ?thesis . qed lemma content_dvd [simp]: "[:content p:] dvd p" by (subst const_poly_dvd_iff_dvd_content) simp_all lemma content_dvd_coeff [simp]: "content p dvd coeff p n" proof (cases "p = 0") case True then show ?thesis by simp next case False then show ?thesis by (cases "n \ degree p") (auto simp add: content_def not_le coeff_eq_0 coeff_in_coeffs intro: Gcd_fin_dvd) qed lemma content_dvd_coeffs: "c \ set (coeffs p) \ content p dvd c" by (simp add: content_def Gcd_fin_dvd) lemma normalize_content [simp]: "normalize (content p) = content p" by (simp add: content_def) lemma is_unit_content_iff [simp]: "is_unit (content p) \ content p = 1" proof assume "is_unit (content p)" then have "normalize (content p) = 1" by (simp add: is_unit_normalize del: normalize_content) then show "content p = 1" by simp qed auto lemma content_smult [simp]: fixes c :: "'a :: {normalization_semidom_multiplicative, semiring_gcd}" shows "content (smult c p) = normalize c * content p" by (simp add: content_def coeffs_smult Gcd_fin_mult normalize_mult) lemma content_eq_zero_iff [simp]: "content p = 0 \ p = 0" by (auto simp: content_def simp: poly_eq_iff coeffs_def) definition primitive_part :: "'a :: semiring_gcd poly \ 'a poly" where "primitive_part p = map_poly (\x. x div content p) p" lemma primitive_part_0 [simp]: "primitive_part 0 = 0" by (simp add: primitive_part_def) lemma content_times_primitive_part [simp]: "smult (content p) (primitive_part p) = p" for p :: "'a :: semiring_gcd poly" proof (cases "p = 0") case True then show ?thesis by simp next case False then show ?thesis unfolding primitive_part_def by (auto simp: smult_conv_map_poly map_poly_map_poly o_def content_dvd_coeffs intro: map_poly_idI) qed lemma primitive_part_eq_0_iff [simp]: "primitive_part p = 0 \ p = 0" proof (cases "p = 0") case True then show ?thesis by simp next case False then have "primitive_part p = map_poly (\x. x div content p) p" by (simp add: primitive_part_def) also from False have "\ = 0 \ p = 0" by (intro map_poly_eq_0_iff) (auto simp: dvd_div_eq_0_iff content_dvd_coeffs) finally show ?thesis using False by simp qed lemma content_primitive_part [simp]: fixes p :: "'a :: {normalization_semidom_multiplicative, semiring_gcd} poly" assumes "p \ 0" shows "content (primitive_part p) = 1" proof - have "p = smult (content p) (primitive_part p)" by simp also have "content \ = content (primitive_part p) * content p" by (simp del: content_times_primitive_part add: ac_simps) finally have "1 * content p = content (primitive_part p) * content p" by simp then have "1 * content p div content p = content (primitive_part p) * content p div content p" by simp with assms show ?thesis by simp qed lemma content_decompose: obtains p' :: "'a :: {normalization_semidom_multiplicative, semiring_gcd} poly" where "p = smult (content p) p'" "content p' = 1" proof (cases "p = 0") case True then have "p = smult (content p) 1" "content 1 = 1" by simp_all then show ?thesis .. next case False then have "p = smult (content p) (primitive_part p)" "content (primitive_part p) = 1" by simp_all then show ?thesis .. qed lemma content_dvd_contentI [intro]: "p dvd q \ content p dvd content q" using const_poly_dvd_iff_dvd_content content_dvd dvd_trans by blast lemma primitive_part_const_poly [simp]: "primitive_part [:x:] = [:unit_factor x:]" by (simp add: primitive_part_def map_poly_pCons) lemma primitive_part_prim: "content p = 1 \ primitive_part p = p" by (auto simp: primitive_part_def) lemma degree_primitive_part [simp]: "degree (primitive_part p) = degree p" proof (cases "p = 0") case True then show ?thesis by simp next case False have "p = smult (content p) (primitive_part p)" by simp also from False have "degree \ = degree (primitive_part p)" by (subst degree_smult_eq) simp_all finally show ?thesis .. qed lemma smult_content_normalize_primitive_part [simp]: fixes p :: "'a :: {normalization_semidom_multiplicative, semiring_gcd, idom_divide} poly" shows "smult (content p) (normalize (primitive_part p)) = normalize p" proof - have "smult (content p) (normalize (primitive_part p)) = normalize ([:content p:] * primitive_part p)" by (subst normalize_mult) (simp_all add: normalize_const_poly) also have "[:content p:] * primitive_part p = p" by simp finally show ?thesis . qed context begin private lemma content_1_mult: fixes f g :: "'a :: {semiring_gcd, factorial_semiring} poly" assumes "content f = 1" "content g = 1" shows "content (f * g) = 1" proof (cases "f * g = 0") case False from assms have "f \ 0" "g \ 0" by auto hence "f * g \ 0" by auto { assume "\is_unit (content (f * g))" with False have "\p. p dvd content (f * g) \ prime p" by (intro prime_divisor_exists) simp_all then obtain p where "p dvd content (f * g)" "prime p" by blast from \p dvd content (f * g)\ have "[:p:] dvd f * g" by (simp add: const_poly_dvd_iff_dvd_content) moreover from \prime p\ have "prime_elem [:p:]" by (simp add: lift_prime_elem_poly) ultimately have "[:p:] dvd f \ [:p:] dvd g" by (simp add: prime_elem_dvd_mult_iff) with assms have "is_unit p" by (simp add: const_poly_dvd_iff_dvd_content) with \prime p\ have False by simp } hence "is_unit (content (f * g))" by blast hence "normalize (content (f * g)) = 1" by (simp add: is_unit_normalize del: normalize_content) thus ?thesis by simp qed (insert assms, auto) lemma content_mult: fixes p q :: "'a :: {factorial_semiring, semiring_gcd, normalization_semidom_multiplicative} poly" shows "content (p * q) = content p * content q" proof (cases "p * q = 0") case False then have "p \ 0" and "q \ 0" by simp_all then have *: "content (primitive_part p * primitive_part q) = 1" by (auto intro: content_1_mult) have "p * q = smult (content p) (primitive_part p) * smult (content q) (primitive_part q)" by simp also have "\ = smult (content p * content q) (primitive_part p * primitive_part q)" by (metis mult.commute mult_smult_right smult_smult) with * show ?thesis by (simp add: normalize_mult) next case True then show ?thesis by auto qed end lemma primitive_part_mult: fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide, normalization_semidom_multiplicative} poly" shows "primitive_part (p * q) = primitive_part p * primitive_part q" proof - have "primitive_part (p * q) = p * q div [:content (p * q):]" by (simp add: primitive_part_def div_const_poly_conv_map_poly) also have "\ = (p div [:content p:]) * (q div [:content q:])" by (subst div_mult_div_if_dvd) (simp_all add: content_mult mult_ac) also have "\ = primitive_part p * primitive_part q" by (simp add: primitive_part_def div_const_poly_conv_map_poly) finally show ?thesis . qed lemma primitive_part_smult: fixes p :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide, normalization_semidom_multiplicative} poly" shows "primitive_part (smult a p) = smult (unit_factor a) (primitive_part p)" proof - have "smult a p = [:a:] * p" by simp also have "primitive_part \ = smult (unit_factor a) (primitive_part p)" by (subst primitive_part_mult) simp_all finally show ?thesis . qed lemma primitive_part_dvd_primitive_partI [intro]: fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide, normalization_semidom_multiplicative} poly" shows "p dvd q \ primitive_part p dvd primitive_part q" by (auto elim!: dvdE simp: primitive_part_mult) lemma content_prod_mset: fixes A :: "'a :: {factorial_semiring, semiring_Gcd, normalization_semidom_multiplicative} poly multiset" shows "content (prod_mset A) = prod_mset (image_mset content A)" by (induction A) (simp_all add: content_mult mult_ac) lemma content_prod_eq_1_iff: fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, normalization_semidom_multiplicative} poly" shows "content (p * q) = 1 \ content p = 1 \ content q = 1" proof safe assume A: "content (p * q) = 1" { fix p q :: "'a poly" assume "content p * content q = 1" hence "1 = content p * content q" by simp hence "content p dvd 1" by (rule dvdI) hence "content p = 1" by simp } note B = this from A B[of p q] B [of q p] show "content p = 1" "content q = 1" by (simp_all add: content_mult mult_ac) qed (auto simp: content_mult) no_notation cCons (infixr "##" 65) end diff --git a/src/HOL/GCD.thy b/src/HOL/GCD.thy --- a/src/HOL/GCD.thy +++ b/src/HOL/GCD.thy @@ -1,2787 +1,2925 @@ (* Title: HOL/GCD.thy Author: Christophe Tabacznyj Author: Lawrence C. Paulson Author: Amine Chaieb Author: Thomas M. Rasmussen Author: Jeremy Avigad Author: Tobias Nipkow This file deals with the functions gcd and lcm. Definitions and lemmas are proved uniformly for the natural numbers and integers. This file combines and revises a number of prior developments. The original theories "GCD" and "Primes" were by Christophe Tabacznyj and Lawrence C. Paulson, based on @{cite davenport92}. They introduced gcd, lcm, and prime for the natural numbers. The original theory "IntPrimes" was by Thomas M. Rasmussen, and extended gcd, lcm, primes to the integers. Amine Chaieb provided another extension of the notions to the integers, and added a number of results to "Primes" and "GCD". IntPrimes also defined and developed the congruence relations on the integers. The notion was extended to the natural numbers by Chaieb. Jeremy Avigad combined all of these, made everything uniform for the natural numbers and the integers, and added a number of new theorems. Tobias Nipkow cleaned up a lot. *) section \Greatest common divisor and least common multiple\ theory GCD imports Groups_List begin subsection \Abstract bounded quasi semilattices as common foundation\ locale bounded_quasi_semilattice = abel_semigroup + fixes top :: 'a ("\<^bold>\") and bot :: 'a ("\<^bold>\") and normalize :: "'a \ 'a" assumes idem_normalize [simp]: "a \<^bold>* a = normalize a" and normalize_left_idem [simp]: "normalize a \<^bold>* b = a \<^bold>* b" and normalize_idem [simp]: "normalize (a \<^bold>* b) = a \<^bold>* b" and normalize_top [simp]: "normalize \<^bold>\ = \<^bold>\" and normalize_bottom [simp]: "normalize \<^bold>\ = \<^bold>\" and top_left_normalize [simp]: "\<^bold>\ \<^bold>* a = normalize a" and bottom_left_bottom [simp]: "\<^bold>\ \<^bold>* a = \<^bold>\" begin lemma left_idem [simp]: "a \<^bold>* (a \<^bold>* b) = a \<^bold>* b" using assoc [of a a b, symmetric] by simp lemma right_idem [simp]: "(a \<^bold>* b) \<^bold>* b = a \<^bold>* b" using left_idem [of b a] by (simp add: ac_simps) lemma comp_fun_idem: "comp_fun_idem f" by standard (simp_all add: fun_eq_iff ac_simps) interpretation comp_fun_idem f by (fact comp_fun_idem) lemma top_right_normalize [simp]: "a \<^bold>* \<^bold>\ = normalize a" using top_left_normalize [of a] by (simp add: ac_simps) lemma bottom_right_bottom [simp]: "a \<^bold>* \<^bold>\ = \<^bold>\" using bottom_left_bottom [of a] by (simp add: ac_simps) lemma normalize_right_idem [simp]: "a \<^bold>* normalize b = a \<^bold>* b" using normalize_left_idem [of b a] by (simp add: ac_simps) end locale bounded_quasi_semilattice_set = bounded_quasi_semilattice begin interpretation comp_fun_idem f by (fact comp_fun_idem) definition F :: "'a set \ 'a" where eq_fold: "F A = (if finite A then Finite_Set.fold f \<^bold>\ A else \<^bold>\)" lemma infinite [simp]: "infinite A \ F A = \<^bold>\" by (simp add: eq_fold) lemma set_eq_fold [code]: "F (set xs) = fold f xs \<^bold>\" by (simp add: eq_fold fold_set_fold) lemma empty [simp]: "F {} = \<^bold>\" by (simp add: eq_fold) lemma insert [simp]: "F (insert a A) = a \<^bold>* F A" by (cases "finite A") (simp_all add: eq_fold) lemma normalize [simp]: "normalize (F A) = F A" by (induct A rule: infinite_finite_induct) simp_all lemma in_idem: assumes "a \ A" shows "a \<^bold>* F A = F A" using assms by (induct A rule: infinite_finite_induct) (auto simp: left_commute [of a]) lemma union: "F (A \ B) = F A \<^bold>* F B" by (induct A rule: infinite_finite_induct) (simp_all add: ac_simps) lemma remove: assumes "a \ A" shows "F A = a \<^bold>* F (A - {a})" proof - from assms obtain B where "A = insert a B" and "a \ B" by (blast dest: mk_disjoint_insert) with assms show ?thesis by simp qed lemma insert_remove: "F (insert a A) = a \<^bold>* F (A - {a})" by (cases "a \ A") (simp_all add: insert_absorb remove) lemma subset: assumes "B \ A" shows "F B \<^bold>* F A = F A" using assms by (simp add: union [symmetric] Un_absorb1) end subsection \Abstract GCD and LCM\ class gcd = zero + one + dvd + fixes gcd :: "'a \ 'a \ 'a" and lcm :: "'a \ 'a \ 'a" class Gcd = gcd + fixes Gcd :: "'a set \ 'a" and Lcm :: "'a set \ 'a" syntax "_GCD1" :: "pttrns \ 'b \ 'b" ("(3GCD _./ _)" [0, 10] 10) "_GCD" :: "pttrn \ 'a set \ 'b \ 'b" ("(3GCD _\_./ _)" [0, 0, 10] 10) "_LCM1" :: "pttrns \ 'b \ 'b" ("(3LCM _./ _)" [0, 10] 10) "_LCM" :: "pttrn \ 'a set \ 'b \ 'b" ("(3LCM _\_./ _)" [0, 0, 10] 10) translations "GCD x y. f" \ "GCD x. GCD y. f" "GCD x. f" \ "CONST Gcd (CONST range (\x. f))" "GCD x\A. f" \ "CONST Gcd ((\x. f) ` A)" "LCM x y. f" \ "LCM x. LCM y. f" "LCM x. f" \ "CONST Lcm (CONST range (\x. f))" "LCM x\A. f" \ "CONST Lcm ((\x. f) ` A)" class semiring_gcd = normalization_semidom + gcd + assumes gcd_dvd1 [iff]: "gcd a b dvd a" and gcd_dvd2 [iff]: "gcd a b dvd b" and gcd_greatest: "c dvd a \ c dvd b \ c dvd gcd a b" and normalize_gcd [simp]: "normalize (gcd a b) = gcd a b" and lcm_gcd: "lcm a b = normalize (a * b div gcd a b)" begin lemma gcd_greatest_iff [simp]: "a dvd gcd b c \ a dvd b \ a dvd c" by (blast intro!: gcd_greatest intro: dvd_trans) lemma gcd_dvdI1: "a dvd c \ gcd a b dvd c" by (rule dvd_trans) (rule gcd_dvd1) lemma gcd_dvdI2: "b dvd c \ gcd a b dvd c" by (rule dvd_trans) (rule gcd_dvd2) lemma dvd_gcdD1: "a dvd gcd b c \ a dvd b" using gcd_dvd1 [of b c] by (blast intro: dvd_trans) lemma dvd_gcdD2: "a dvd gcd b c \ a dvd c" using gcd_dvd2 [of b c] by (blast intro: dvd_trans) lemma gcd_0_left [simp]: "gcd 0 a = normalize a" by (rule associated_eqI) simp_all lemma gcd_0_right [simp]: "gcd a 0 = normalize a" by (rule associated_eqI) simp_all lemma gcd_eq_0_iff [simp]: "gcd a b = 0 \ a = 0 \ b = 0" (is "?P \ ?Q") proof assume ?P then have "0 dvd gcd a b" by simp then have "0 dvd a" and "0 dvd b" by (blast intro: dvd_trans)+ then show ?Q by simp next assume ?Q then show ?P by simp qed lemma unit_factor_gcd: "unit_factor (gcd a b) = (if a = 0 \ b = 0 then 0 else 1)" proof (cases "gcd a b = 0") case True then show ?thesis by simp next case False have "unit_factor (gcd a b) * normalize (gcd a b) = gcd a b" by (rule unit_factor_mult_normalize) then have "unit_factor (gcd a b) * gcd a b = gcd a b" by simp then have "unit_factor (gcd a b) * gcd a b div gcd a b = gcd a b div gcd a b" by simp with False show ?thesis by simp qed lemma is_unit_gcd_iff [simp]: "is_unit (gcd a b) \ gcd a b = 1" by (cases "a = 0 \ b = 0") (auto simp: unit_factor_gcd dest: is_unit_unit_factor) sublocale gcd: abel_semigroup gcd proof fix a b c show "gcd a b = gcd b a" by (rule associated_eqI) simp_all from gcd_dvd1 have "gcd (gcd a b) c dvd a" by (rule dvd_trans) simp moreover from gcd_dvd1 have "gcd (gcd a b) c dvd b" by (rule dvd_trans) simp ultimately have P1: "gcd (gcd a b) c dvd gcd a (gcd b c)" by (auto intro!: gcd_greatest) from gcd_dvd2 have "gcd a (gcd b c) dvd b" by (rule dvd_trans) simp moreover from gcd_dvd2 have "gcd a (gcd b c) dvd c" by (rule dvd_trans) simp ultimately have P2: "gcd a (gcd b c) dvd gcd (gcd a b) c" by (auto intro!: gcd_greatest) from P1 P2 show "gcd (gcd a b) c = gcd a (gcd b c)" by (rule associated_eqI) simp_all qed sublocale gcd: bounded_quasi_semilattice gcd 0 1 normalize proof show "gcd a a = normalize a" for a proof - have "a dvd gcd a a" by (rule gcd_greatest) simp_all then show ?thesis by (auto intro: associated_eqI) qed show "gcd (normalize a) b = gcd a b" for a b using gcd_dvd1 [of "normalize a" b] by (auto intro: associated_eqI) show "gcd 1 a = 1" for a by (rule associated_eqI) simp_all qed simp_all lemma gcd_self: "gcd a a = normalize a" by (fact gcd.idem_normalize) lemma gcd_left_idem: "gcd a (gcd a b) = gcd a b" by (fact gcd.left_idem) lemma gcd_right_idem: "gcd (gcd a b) b = gcd a b" by (fact gcd.right_idem) lemma gcdI: assumes "c dvd a" and "c dvd b" and greatest: "\d. d dvd a \ d dvd b \ d dvd c" and "normalize c = c" shows "c = gcd a b" by (rule associated_eqI) (auto simp: assms intro: gcd_greatest) lemma gcd_unique: "d dvd a \ d dvd b \ normalize d = d \ (\e. e dvd a \ e dvd b \ e dvd d) \ d = gcd a b" by rule (auto intro: gcdI simp: gcd_greatest) lemma gcd_dvd_prod: "gcd a b dvd k * b" using mult_dvd_mono [of 1] by auto lemma gcd_proj2_if_dvd: "b dvd a \ gcd a b = normalize b" by (rule gcdI [symmetric]) simp_all lemma gcd_proj1_if_dvd: "a dvd b \ gcd a b = normalize a" by (rule gcdI [symmetric]) simp_all lemma gcd_proj1_iff: "gcd m n = normalize m \ m dvd n" proof assume *: "gcd m n = normalize m" show "m dvd n" proof (cases "m = 0") case True with * show ?thesis by simp next case [simp]: False from * have **: "m = gcd m n * unit_factor m" by (simp add: unit_eq_div2) show ?thesis by (subst **) (simp add: mult_unit_dvd_iff) qed next assume "m dvd n" then show "gcd m n = normalize m" by (rule gcd_proj1_if_dvd) qed lemma gcd_proj2_iff: "gcd m n = normalize n \ n dvd m" using gcd_proj1_iff [of n m] by (simp add: ac_simps) lemma gcd_mult_left: "gcd (c * a) (c * b) = normalize (c * gcd a b)" proof (cases "c = 0") case True then show ?thesis by simp next case False then have *: "c * gcd a b dvd gcd (c * a) (c * b)" by (auto intro: gcd_greatest) moreover from False * have "gcd (c * a) (c * b) dvd c * gcd a b" by (metis div_dvd_iff_mult dvd_mult_left gcd_dvd1 gcd_dvd2 gcd_greatest mult_commute) ultimately have "normalize (gcd (c * a) (c * b)) = normalize (c * gcd a b)" by (auto intro: associated_eqI) then show ?thesis by (simp add: normalize_mult) qed lemma gcd_mult_right: "gcd (a * c) (b * c) = normalize (gcd b a * c)" using gcd_mult_left [of c a b] by (simp add: ac_simps) lemma dvd_lcm1 [iff]: "a dvd lcm a b" by (metis div_mult_swap dvd_mult2 dvd_normalize_iff dvd_refl gcd_dvd2 lcm_gcd) lemma dvd_lcm2 [iff]: "b dvd lcm a b" by (metis dvd_div_mult dvd_mult dvd_normalize_iff dvd_refl gcd_dvd1 lcm_gcd) lemma dvd_lcmI1: "a dvd b \ a dvd lcm b c" by (rule dvd_trans) (assumption, blast) lemma dvd_lcmI2: "a dvd c \ a dvd lcm b c" by (rule dvd_trans) (assumption, blast) lemma lcm_dvdD1: "lcm a b dvd c \ a dvd c" using dvd_lcm1 [of a b] by (blast intro: dvd_trans) lemma lcm_dvdD2: "lcm a b dvd c \ b dvd c" using dvd_lcm2 [of a b] by (blast intro: dvd_trans) lemma lcm_least: assumes "a dvd c" and "b dvd c" shows "lcm a b dvd c" proof (cases "c = 0") case True then show ?thesis by simp next case False then have *: "is_unit (unit_factor c)" by simp show ?thesis proof (cases "gcd a b = 0") case True with assms show ?thesis by simp next case False have "a * b dvd normalize (c * gcd a b)" using assms by (subst gcd_mult_left [symmetric]) (auto intro!: gcd_greatest simp: mult_ac) with False have "(a * b div gcd a b) dvd c" by (subst div_dvd_iff_mult) auto thus ?thesis by (simp add: lcm_gcd) qed qed lemma lcm_least_iff [simp]: "lcm a b dvd c \ a dvd c \ b dvd c" by (blast intro!: lcm_least intro: dvd_trans) lemma normalize_lcm [simp]: "normalize (lcm a b) = lcm a b" by (simp add: lcm_gcd dvd_normalize_div) lemma lcm_0_left [simp]: "lcm 0 a = 0" by (simp add: lcm_gcd) lemma lcm_0_right [simp]: "lcm a 0 = 0" by (simp add: lcm_gcd) lemma lcm_eq_0_iff: "lcm a b = 0 \ a = 0 \ b = 0" (is "?P \ ?Q") proof assume ?P then have "0 dvd lcm a b" by simp also have "lcm a b dvd (a * b)" by simp finally show ?Q by auto next assume ?Q then show ?P by auto qed lemma zero_eq_lcm_iff: "0 = lcm a b \ a = 0 \ b = 0" using lcm_eq_0_iff[of a b] by auto lemma lcm_eq_1_iff [simp]: "lcm a b = 1 \ is_unit a \ is_unit b" by (auto intro: associated_eqI) lemma unit_factor_lcm: "unit_factor (lcm a b) = (if a = 0 \ b = 0 then 0 else 1)" using lcm_eq_0_iff[of a b] by (cases "lcm a b = 0") (auto simp: lcm_gcd) sublocale lcm: abel_semigroup lcm proof fix a b c show "lcm a b = lcm b a" by (simp add: lcm_gcd ac_simps normalize_mult dvd_normalize_div) have "lcm (lcm a b) c dvd lcm a (lcm b c)" and "lcm a (lcm b c) dvd lcm (lcm a b) c" by (auto intro: lcm_least dvd_trans [of b "lcm b c" "lcm a (lcm b c)"] dvd_trans [of c "lcm b c" "lcm a (lcm b c)"] dvd_trans [of a "lcm a b" "lcm (lcm a b) c"] dvd_trans [of b "lcm a b" "lcm (lcm a b) c"]) then show "lcm (lcm a b) c = lcm a (lcm b c)" by (rule associated_eqI) simp_all qed sublocale lcm: bounded_quasi_semilattice lcm 1 0 normalize proof show "lcm a a = normalize a" for a proof - have "lcm a a dvd a" by (rule lcm_least) simp_all then show ?thesis by (auto intro: associated_eqI) qed show "lcm (normalize a) b = lcm a b" for a b using dvd_lcm1 [of "normalize a" b] unfolding normalize_dvd_iff by (auto intro: associated_eqI) show "lcm 1 a = normalize a" for a by (rule associated_eqI) simp_all qed simp_all lemma lcm_self: "lcm a a = normalize a" by (fact lcm.idem_normalize) lemma lcm_left_idem: "lcm a (lcm a b) = lcm a b" by (fact lcm.left_idem) lemma lcm_right_idem: "lcm (lcm a b) b = lcm a b" by (fact lcm.right_idem) lemma gcd_lcm: assumes "a \ 0" and "b \ 0" shows "gcd a b = normalize (a * b div lcm a b)" proof - from assms have [simp]: "a * b div gcd a b \ 0" by (subst dvd_div_eq_0_iff) auto let ?u = "unit_factor (a * b div gcd a b)" have "gcd a b * normalize (a * b div gcd a b) = gcd a b * ((a * b div gcd a b) * (1 div ?u))" by simp also have "\ = a * b div ?u" by (subst mult.assoc [symmetric]) auto also have "\ dvd a * b" by (subst div_unit_dvd_iff) auto finally have "gcd a b dvd ((a * b) div lcm a b)" by (intro dvd_mult_imp_div) (auto simp: lcm_gcd) moreover have "a * b div lcm a b dvd a" and "a * b div lcm a b dvd b" using assms by (subst div_dvd_iff_mult; simp add: lcm_eq_0_iff mult.commute[of b "lcm a b"])+ ultimately have "normalize (gcd a b) = normalize (a * b div lcm a b)" apply - apply (rule associated_eqI) using assms apply (auto simp: div_dvd_iff_mult zero_eq_lcm_iff) done thus ?thesis by simp qed lemma lcm_1_left: "lcm 1 a = normalize a" by (fact lcm.top_left_normalize) lemma lcm_1_right: "lcm a 1 = normalize a" by (fact lcm.top_right_normalize) lemma lcm_mult_left: "lcm (c * a) (c * b) = normalize (c * lcm a b)" proof (cases "c = 0") case True then show ?thesis by simp next case False then have *: "lcm (c * a) (c * b) dvd c * lcm a b" by (auto intro: lcm_least) moreover have "lcm a b dvd lcm (c * a) (c * b) div c" by (intro lcm_least) (auto intro!: dvd_mult_imp_div simp: mult_ac) hence "c * lcm a b dvd lcm (c * a) (c * b)" using False by (subst (asm) dvd_div_iff_mult) (auto simp: mult_ac intro: dvd_lcmI1) ultimately have "normalize (lcm (c * a) (c * b)) = normalize (c * lcm a b)" by (auto intro: associated_eqI) then show ?thesis by (simp add: normalize_mult) qed lemma lcm_mult_right: "lcm (a * c) (b * c) = normalize (lcm b a * c)" using lcm_mult_left [of c a b] by (simp add: ac_simps) lemma lcm_mult_unit1: "is_unit a \ lcm (b * a) c = lcm b c" by (rule associated_eqI) (simp_all add: mult_unit_dvd_iff dvd_lcmI1) lemma lcm_mult_unit2: "is_unit a \ lcm b (c * a) = lcm b c" using lcm_mult_unit1 [of a c b] by (simp add: ac_simps) lemma lcm_div_unit1: "is_unit a \ lcm (b div a) c = lcm b c" by (erule is_unitE [of _ b]) (simp add: lcm_mult_unit1) lemma lcm_div_unit2: "is_unit a \ lcm b (c div a) = lcm b c" by (erule is_unitE [of _ c]) (simp add: lcm_mult_unit2) lemma normalize_lcm_left: "lcm (normalize a) b = lcm a b" by (fact lcm.normalize_left_idem) lemma normalize_lcm_right: "lcm a (normalize b) = lcm a b" by (fact lcm.normalize_right_idem) lemma comp_fun_idem_gcd: "comp_fun_idem gcd" by standard (simp_all add: fun_eq_iff ac_simps) lemma comp_fun_idem_lcm: "comp_fun_idem lcm" by standard (simp_all add: fun_eq_iff ac_simps) lemma gcd_dvd_antisym: "gcd a b dvd gcd c d \ gcd c d dvd gcd a b \ gcd a b = gcd c d" proof (rule gcdI) assume *: "gcd a b dvd gcd c d" and **: "gcd c d dvd gcd a b" have "gcd c d dvd c" by simp with * show "gcd a b dvd c" by (rule dvd_trans) have "gcd c d dvd d" by simp with * show "gcd a b dvd d" by (rule dvd_trans) show "normalize (gcd a b) = gcd a b" by simp fix l assume "l dvd c" and "l dvd d" then have "l dvd gcd c d" by (rule gcd_greatest) from this and ** show "l dvd gcd a b" by (rule dvd_trans) qed declare unit_factor_lcm [simp] lemma lcmI: assumes "a dvd c" and "b dvd c" and "\d. a dvd d \ b dvd d \ c dvd d" and "normalize c = c" shows "c = lcm a b" by (rule associated_eqI) (auto simp: assms intro: lcm_least) lemma gcd_dvd_lcm [simp]: "gcd a b dvd lcm a b" using gcd_dvd2 by (rule dvd_lcmI2) lemmas lcm_0 = lcm_0_right lemma lcm_unique: "a dvd d \ b dvd d \ normalize d = d \ (\e. a dvd e \ b dvd e \ d dvd e) \ d = lcm a b" by rule (auto intro: lcmI simp: lcm_least lcm_eq_0_iff) lemma lcm_proj1_if_dvd: assumes "b dvd a" shows "lcm a b = normalize a" proof - have "normalize (lcm a b) = normalize a" by (rule associatedI) (use assms in auto) thus ?thesis by simp qed lemma lcm_proj2_if_dvd: "a dvd b \ lcm a b = normalize b" using lcm_proj1_if_dvd [of a b] by (simp add: ac_simps) lemma lcm_proj1_iff: "lcm m n = normalize m \ n dvd m" proof assume *: "lcm m n = normalize m" show "n dvd m" proof (cases "m = 0") case True then show ?thesis by simp next case [simp]: False from * have **: "m = lcm m n * unit_factor m" by (simp add: unit_eq_div2) show ?thesis by (subst **) simp qed next assume "n dvd m" then show "lcm m n = normalize m" by (rule lcm_proj1_if_dvd) qed lemma lcm_proj2_iff: "lcm m n = normalize n \ m dvd n" using lcm_proj1_iff [of n m] by (simp add: ac_simps) lemma gcd_mono: "a dvd c \ b dvd d \ gcd a b dvd gcd c d" by (simp add: gcd_dvdI1 gcd_dvdI2) lemma lcm_mono: "a dvd c \ b dvd d \ lcm a b dvd lcm c d" by (simp add: dvd_lcmI1 dvd_lcmI2) lemma dvd_productE: assumes "p dvd a * b" obtains x y where "p = x * y" "x dvd a" "y dvd b" proof (cases "a = 0") case True thus ?thesis by (intro that[of p 1]) simp_all next case False define x y where "x = gcd a p" and "y = p div x" have "p = x * y" by (simp add: x_def y_def) moreover have "x dvd a" by (simp add: x_def) moreover from assms have "p dvd gcd (b * a) (b * p)" by (intro gcd_greatest) (simp_all add: mult.commute) hence "p dvd b * gcd a p" by (subst (asm) gcd_mult_left) auto with False have "y dvd b" by (simp add: x_def y_def div_dvd_iff_mult assms) ultimately show ?thesis by (rule that) qed lemma gcd_mult_unit1: assumes "is_unit a" shows "gcd (b * a) c = gcd b c" proof - have "gcd (b * a) c dvd b" using assms dvd_mult_unit_iff by blast then show ?thesis by (rule gcdI) simp_all qed lemma gcd_mult_unit2: "is_unit a \ gcd b (c * a) = gcd b c" using gcd.commute gcd_mult_unit1 by auto lemma gcd_div_unit1: "is_unit a \ gcd (b div a) c = gcd b c" by (erule is_unitE [of _ b]) (simp add: gcd_mult_unit1) lemma gcd_div_unit2: "is_unit a \ gcd b (c div a) = gcd b c" by (erule is_unitE [of _ c]) (simp add: gcd_mult_unit2) lemma normalize_gcd_left: "gcd (normalize a) b = gcd a b" by (fact gcd.normalize_left_idem) lemma normalize_gcd_right: "gcd a (normalize b) = gcd a b" by (fact gcd.normalize_right_idem) lemma gcd_add1 [simp]: "gcd (m + n) n = gcd m n" by (rule gcdI [symmetric]) (simp_all add: dvd_add_left_iff) lemma gcd_add2 [simp]: "gcd m (m + n) = gcd m n" using gcd_add1 [of n m] by (simp add: ac_simps) lemma gcd_add_mult: "gcd m (k * m + n) = gcd m n" by (rule gcdI [symmetric]) (simp_all add: dvd_add_right_iff) end class ring_gcd = comm_ring_1 + semiring_gcd begin lemma gcd_neg1 [simp]: "gcd (-a) b = gcd a b" by (rule sym, rule gcdI) (simp_all add: gcd_greatest) lemma gcd_neg2 [simp]: "gcd a (-b) = gcd a b" by (rule sym, rule gcdI) (simp_all add: gcd_greatest) lemma gcd_neg_numeral_1 [simp]: "gcd (- numeral n) a = gcd (numeral n) a" by (fact gcd_neg1) lemma gcd_neg_numeral_2 [simp]: "gcd a (- numeral n) = gcd a (numeral n)" by (fact gcd_neg2) lemma gcd_diff1: "gcd (m - n) n = gcd m n" by (subst diff_conv_add_uminus, subst gcd_neg2[symmetric], subst gcd_add1, simp) lemma gcd_diff2: "gcd (n - m) n = gcd m n" by (subst gcd_neg1[symmetric]) (simp only: minus_diff_eq gcd_diff1) lemma lcm_neg1 [simp]: "lcm (-a) b = lcm a b" by (rule sym, rule lcmI) (simp_all add: lcm_least lcm_eq_0_iff) lemma lcm_neg2 [simp]: "lcm a (-b) = lcm a b" by (rule sym, rule lcmI) (simp_all add: lcm_least lcm_eq_0_iff) lemma lcm_neg_numeral_1 [simp]: "lcm (- numeral n) a = lcm (numeral n) a" by (fact lcm_neg1) lemma lcm_neg_numeral_2 [simp]: "lcm a (- numeral n) = lcm a (numeral n)" by (fact lcm_neg2) end class semiring_Gcd = semiring_gcd + Gcd + assumes Gcd_dvd: "a \ A \ Gcd A dvd a" and Gcd_greatest: "(\b. b \ A \ a dvd b) \ a dvd Gcd A" and normalize_Gcd [simp]: "normalize (Gcd A) = Gcd A" assumes dvd_Lcm: "a \ A \ a dvd Lcm A" and Lcm_least: "(\b. b \ A \ b dvd a) \ Lcm A dvd a" and normalize_Lcm [simp]: "normalize (Lcm A) = Lcm A" begin lemma Lcm_Gcd: "Lcm A = Gcd {b. \a\A. a dvd b}" by (rule associated_eqI) (auto intro: Gcd_dvd dvd_Lcm Gcd_greatest Lcm_least) lemma Gcd_Lcm: "Gcd A = Lcm {b. \a\A. b dvd a}" by (rule associated_eqI) (auto intro: Gcd_dvd dvd_Lcm Gcd_greatest Lcm_least) lemma Gcd_empty [simp]: "Gcd {} = 0" by (rule dvd_0_left, rule Gcd_greatest) simp lemma Lcm_empty [simp]: "Lcm {} = 1" by (auto intro: associated_eqI Lcm_least) lemma Gcd_insert [simp]: "Gcd (insert a A) = gcd a (Gcd A)" proof - have "Gcd (insert a A) dvd gcd a (Gcd A)" by (auto intro: Gcd_dvd Gcd_greatest) moreover have "gcd a (Gcd A) dvd Gcd (insert a A)" proof (rule Gcd_greatest) fix b assume "b \ insert a A" then show "gcd a (Gcd A) dvd b" proof assume "b = a" then show ?thesis by simp next assume "b \ A" then have "Gcd A dvd b" by (rule Gcd_dvd) moreover have "gcd a (Gcd A) dvd Gcd A" by simp ultimately show ?thesis by (blast intro: dvd_trans) qed qed ultimately show ?thesis by (auto intro: associated_eqI) qed lemma Lcm_insert [simp]: "Lcm (insert a A) = lcm a (Lcm A)" proof (rule sym) have "lcm a (Lcm A) dvd Lcm (insert a A)" by (auto intro: dvd_Lcm Lcm_least) moreover have "Lcm (insert a A) dvd lcm a (Lcm A)" proof (rule Lcm_least) fix b assume "b \ insert a A" then show "b dvd lcm a (Lcm A)" proof assume "b = a" then show ?thesis by simp next assume "b \ A" then have "b dvd Lcm A" by (rule dvd_Lcm) moreover have "Lcm A dvd lcm a (Lcm A)" by simp ultimately show ?thesis by (blast intro: dvd_trans) qed qed ultimately show "lcm a (Lcm A) = Lcm (insert a A)" by (rule associated_eqI) (simp_all add: lcm_eq_0_iff) qed lemma LcmI: assumes "\a. a \ A \ a dvd b" and "\c. (\a. a \ A \ a dvd c) \ b dvd c" and "normalize b = b" shows "b = Lcm A" by (rule associated_eqI) (auto simp: assms dvd_Lcm intro: Lcm_least) lemma Lcm_subset: "A \ B \ Lcm A dvd Lcm B" by (blast intro: Lcm_least dvd_Lcm) lemma Lcm_Un: "Lcm (A \ B) = lcm (Lcm A) (Lcm B)" proof - have "\d. \Lcm A dvd d; Lcm B dvd d\ \ Lcm (A \ B) dvd d" by (meson UnE Lcm_least dvd_Lcm dvd_trans) then show ?thesis by (meson Lcm_subset lcm_unique normalize_Lcm sup.cobounded1 sup.cobounded2) qed lemma Gcd_0_iff [simp]: "Gcd A = 0 \ A \ {0}" (is "?P \ ?Q") proof assume ?P show ?Q proof fix a assume "a \ A" then have "Gcd A dvd a" by (rule Gcd_dvd) with \?P\ have "a = 0" by simp then show "a \ {0}" by simp qed next assume ?Q have "0 dvd Gcd A" proof (rule Gcd_greatest) fix a assume "a \ A" with \?Q\ have "a = 0" by auto then show "0 dvd a" by simp qed then show ?P by simp qed lemma Lcm_1_iff [simp]: "Lcm A = 1 \ (\a\A. is_unit a)" (is "?P \ ?Q") proof assume ?P show ?Q proof fix a assume "a \ A" then have "a dvd Lcm A" by (rule dvd_Lcm) with \?P\ show "is_unit a" by simp qed next assume ?Q then have "is_unit (Lcm A)" by (blast intro: Lcm_least) then have "normalize (Lcm A) = 1" by (rule is_unit_normalize) then show ?P by simp qed lemma unit_factor_Lcm: "unit_factor (Lcm A) = (if Lcm A = 0 then 0 else 1)" proof (cases "Lcm A = 0") case True then show ?thesis by simp next case False with unit_factor_normalize have "unit_factor (normalize (Lcm A)) = 1" by blast with False show ?thesis by simp qed lemma unit_factor_Gcd: "unit_factor (Gcd A) = (if Gcd A = 0 then 0 else 1)" by (simp add: Gcd_Lcm unit_factor_Lcm) lemma GcdI: assumes "\a. a \ A \ b dvd a" and "\c. (\a. a \ A \ c dvd a) \ c dvd b" and "normalize b = b" shows "b = Gcd A" by (rule associated_eqI) (auto simp: assms Gcd_dvd intro: Gcd_greatest) lemma Gcd_eq_1_I: assumes "is_unit a" and "a \ A" shows "Gcd A = 1" proof - from assms have "is_unit (Gcd A)" by (blast intro: Gcd_dvd dvd_unit_imp_unit) then have "normalize (Gcd A) = 1" by (rule is_unit_normalize) then show ?thesis by simp qed lemma Lcm_eq_0_I: assumes "0 \ A" shows "Lcm A = 0" proof - from assms have "0 dvd Lcm A" by (rule dvd_Lcm) then show ?thesis by simp qed lemma Gcd_UNIV [simp]: "Gcd UNIV = 1" using dvd_refl by (rule Gcd_eq_1_I) simp lemma Lcm_UNIV [simp]: "Lcm UNIV = 0" by (rule Lcm_eq_0_I) simp lemma Lcm_0_iff: assumes "finite A" shows "Lcm A = 0 \ 0 \ A" proof (cases "A = {}") case True then show ?thesis by simp next case False with assms show ?thesis by (induct A rule: finite_ne_induct) (auto simp: lcm_eq_0_iff) qed lemma Gcd_image_normalize [simp]: "Gcd (normalize ` A) = Gcd A" proof - have "Gcd (normalize ` A) dvd a" if "a \ A" for a proof - from that obtain B where "A = insert a B" by blast moreover have "gcd (normalize a) (Gcd (normalize ` B)) dvd normalize a" by (rule gcd_dvd1) ultimately show "Gcd (normalize ` A) dvd a" by simp qed then have "Gcd (normalize ` A) dvd Gcd A" and "Gcd A dvd Gcd (normalize ` A)" by (auto intro!: Gcd_greatest intro: Gcd_dvd) then show ?thesis by (auto intro: associated_eqI) qed lemma Gcd_eqI: assumes "normalize a = a" assumes "\b. b \ A \ a dvd b" and "\c. (\b. b \ A \ c dvd b) \ c dvd a" shows "Gcd A = a" using assms by (blast intro: associated_eqI Gcd_greatest Gcd_dvd normalize_Gcd) lemma dvd_GcdD: "x dvd Gcd A \ y \ A \ x dvd y" using Gcd_dvd dvd_trans by blast lemma dvd_Gcd_iff: "x dvd Gcd A \ (\y\A. x dvd y)" by (blast dest: dvd_GcdD intro: Gcd_greatest) lemma Gcd_mult: "Gcd ((*) c ` A) = normalize (c * Gcd A)" proof (cases "c = 0") case True then show ?thesis by auto next case [simp]: False have "Gcd ((*) c ` A) div c dvd Gcd A" by (intro Gcd_greatest, subst div_dvd_iff_mult) (auto intro!: Gcd_greatest Gcd_dvd simp: mult.commute[of _ c]) then have "Gcd ((*) c ` A) dvd c * Gcd A" by (subst (asm) div_dvd_iff_mult) (auto intro: Gcd_greatest simp: mult_ac) moreover have "c * Gcd A dvd Gcd ((*) c ` A)" by (intro Gcd_greatest) (auto intro: mult_dvd_mono Gcd_dvd) ultimately have "normalize (Gcd ((*) c ` A)) = normalize (c * Gcd A)" by (rule associatedI) then show ?thesis by simp qed lemma Lcm_eqI: assumes "normalize a = a" and "\b. b \ A \ b dvd a" and "\c. (\b. b \ A \ b dvd c) \ a dvd c" shows "Lcm A = a" using assms by (blast intro: associated_eqI Lcm_least dvd_Lcm normalize_Lcm) lemma Lcm_dvdD: "Lcm A dvd x \ y \ A \ y dvd x" using dvd_Lcm dvd_trans by blast lemma Lcm_dvd_iff: "Lcm A dvd x \ (\y\A. y dvd x)" by (blast dest: Lcm_dvdD intro: Lcm_least) lemma Lcm_mult: assumes "A \ {}" shows "Lcm ((*) c ` A) = normalize (c * Lcm A)" proof (cases "c = 0") case True with assms have "(*) c ` A = {0}" by auto with True show ?thesis by auto next case [simp]: False from assms obtain x where x: "x \ A" by blast have "c dvd c * x" by simp also from x have "c * x dvd Lcm ((*) c ` A)" by (intro dvd_Lcm) auto finally have dvd: "c dvd Lcm ((*) c ` A)" . moreover have "Lcm A dvd Lcm ((*) c ` A) div c" by (intro Lcm_least dvd_mult_imp_div) (auto intro!: Lcm_least dvd_Lcm simp: mult.commute[of _ c]) ultimately have "c * Lcm A dvd Lcm ((*) c ` A)" by auto moreover have "Lcm ((*) c ` A) dvd c * Lcm A" by (intro Lcm_least) (auto intro: mult_dvd_mono dvd_Lcm) ultimately have "normalize (c * Lcm A) = normalize (Lcm ((*) c ` A))" by (rule associatedI) then show ?thesis by simp qed lemma Lcm_no_units: "Lcm A = Lcm (A - {a. is_unit a})" proof - have "(A - {a. is_unit a}) \ {a\A. is_unit a} = A" by blast then have "Lcm A = lcm (Lcm (A - {a. is_unit a})) (Lcm {a\A. is_unit a})" by (simp add: Lcm_Un [symmetric]) also have "Lcm {a\A. is_unit a} = 1" by simp finally show ?thesis by simp qed lemma Lcm_0_iff': "Lcm A = 0 \ (\l. l \ 0 \ (\a\A. a dvd l))" by (metis Lcm_least dvd_0_left dvd_Lcm) lemma Lcm_no_multiple: "(\m. m \ 0 \ (\a\A. \ a dvd m)) \ Lcm A = 0" by (auto simp: Lcm_0_iff') lemma Lcm_singleton [simp]: "Lcm {a} = normalize a" by simp lemma Lcm_2 [simp]: "Lcm {a, b} = lcm a b" by simp lemma Gcd_1: "1 \ A \ Gcd A = 1" by (auto intro!: Gcd_eq_1_I) lemma Gcd_singleton [simp]: "Gcd {a} = normalize a" by simp lemma Gcd_2 [simp]: "Gcd {a, b} = gcd a b" by simp lemma Gcd_mono: assumes "\x. x \ A \ f x dvd g x" shows "(GCD x\A. f x) dvd (GCD x\A. g x)" proof (intro Gcd_greatest, safe) fix x assume "x \ A" hence "(GCD x\A. f x) dvd f x" by (intro Gcd_dvd) auto also have "f x dvd g x" using \x \ A\ assms by blast finally show "(GCD x\A. f x) dvd \" . qed lemma Lcm_mono: assumes "\x. x \ A \ f x dvd g x" shows "(LCM x\A. f x) dvd (LCM x\A. g x)" proof (intro Lcm_least, safe) fix x assume "x \ A" hence "f x dvd g x" by (rule assms) also have "g x dvd (LCM x\A. g x)" using \x \ A\ by (intro dvd_Lcm) auto finally show "f x dvd \" . qed end subsection \An aside: GCD and LCM on finite sets for incomplete gcd rings\ context semiring_gcd begin sublocale Gcd_fin: bounded_quasi_semilattice_set gcd 0 1 normalize defines Gcd_fin ("Gcd\<^sub>f\<^sub>i\<^sub>n") = "Gcd_fin.F :: 'a set \ 'a" .. abbreviation gcd_list :: "'a list \ 'a" where "gcd_list xs \ Gcd\<^sub>f\<^sub>i\<^sub>n (set xs)" sublocale Lcm_fin: bounded_quasi_semilattice_set lcm 1 0 normalize defines Lcm_fin ("Lcm\<^sub>f\<^sub>i\<^sub>n") = Lcm_fin.F .. abbreviation lcm_list :: "'a list \ 'a" where "lcm_list xs \ Lcm\<^sub>f\<^sub>i\<^sub>n (set xs)" lemma Gcd_fin_dvd: "a \ A \ Gcd\<^sub>f\<^sub>i\<^sub>n A dvd a" by (induct A rule: infinite_finite_induct) (auto intro: dvd_trans) lemma dvd_Lcm_fin: "a \ A \ a dvd Lcm\<^sub>f\<^sub>i\<^sub>n A" by (induct A rule: infinite_finite_induct) (auto intro: dvd_trans) lemma Gcd_fin_greatest: "a dvd Gcd\<^sub>f\<^sub>i\<^sub>n A" if "finite A" and "\b. b \ A \ a dvd b" using that by (induct A) simp_all lemma Lcm_fin_least: "Lcm\<^sub>f\<^sub>i\<^sub>n A dvd a" if "finite A" and "\b. b \ A \ b dvd a" using that by (induct A) simp_all lemma gcd_list_greatest: "a dvd gcd_list bs" if "\b. b \ set bs \ a dvd b" by (rule Gcd_fin_greatest) (simp_all add: that) lemma lcm_list_least: "lcm_list bs dvd a" if "\b. b \ set bs \ b dvd a" by (rule Lcm_fin_least) (simp_all add: that) lemma dvd_Gcd_fin_iff: "b dvd Gcd\<^sub>f\<^sub>i\<^sub>n A \ (\a\A. b dvd a)" if "finite A" using that by (auto intro: Gcd_fin_greatest Gcd_fin_dvd dvd_trans [of b "Gcd\<^sub>f\<^sub>i\<^sub>n A"]) lemma dvd_gcd_list_iff: "b dvd gcd_list xs \ (\a\set xs. b dvd a)" by (simp add: dvd_Gcd_fin_iff) lemma Lcm_fin_dvd_iff: "Lcm\<^sub>f\<^sub>i\<^sub>n A dvd b \ (\a\A. a dvd b)" if "finite A" using that by (auto intro: Lcm_fin_least dvd_Lcm_fin dvd_trans [of _ "Lcm\<^sub>f\<^sub>i\<^sub>n A" b]) lemma lcm_list_dvd_iff: "lcm_list xs dvd b \ (\a\set xs. a dvd b)" by (simp add: Lcm_fin_dvd_iff) lemma Gcd_fin_mult: "Gcd\<^sub>f\<^sub>i\<^sub>n (image (times b) A) = normalize (b * Gcd\<^sub>f\<^sub>i\<^sub>n A)" if "finite A" using that by induction (auto simp: gcd_mult_left) lemma Lcm_fin_mult: "Lcm\<^sub>f\<^sub>i\<^sub>n (image (times b) A) = normalize (b * Lcm\<^sub>f\<^sub>i\<^sub>n A)" if "A \ {}" proof (cases "b = 0") case True moreover from that have "times 0 ` A = {0}" by auto ultimately show ?thesis by simp next case False show ?thesis proof (cases "finite A") case False moreover have "inj_on (times b) A" using \b \ 0\ by (rule inj_on_mult) ultimately have "infinite (times b ` A)" by (simp add: finite_image_iff) with False show ?thesis by simp next case True then show ?thesis using that by (induct A rule: finite_ne_induct) (auto simp: lcm_mult_left) qed qed lemma unit_factor_Gcd_fin: "unit_factor (Gcd\<^sub>f\<^sub>i\<^sub>n A) = of_bool (Gcd\<^sub>f\<^sub>i\<^sub>n A \ 0)" by (rule normalize_idem_imp_unit_factor_eq) simp lemma unit_factor_Lcm_fin: "unit_factor (Lcm\<^sub>f\<^sub>i\<^sub>n A) = of_bool (Lcm\<^sub>f\<^sub>i\<^sub>n A \ 0)" by (rule normalize_idem_imp_unit_factor_eq) simp lemma is_unit_Gcd_fin_iff [simp]: "is_unit (Gcd\<^sub>f\<^sub>i\<^sub>n A) \ Gcd\<^sub>f\<^sub>i\<^sub>n A = 1" by (rule normalize_idem_imp_is_unit_iff) simp lemma is_unit_Lcm_fin_iff [simp]: "is_unit (Lcm\<^sub>f\<^sub>i\<^sub>n A) \ Lcm\<^sub>f\<^sub>i\<^sub>n A = 1" by (rule normalize_idem_imp_is_unit_iff) simp lemma Gcd_fin_0_iff: "Gcd\<^sub>f\<^sub>i\<^sub>n A = 0 \ A \ {0} \ finite A" by (induct A rule: infinite_finite_induct) simp_all lemma Lcm_fin_0_iff: "Lcm\<^sub>f\<^sub>i\<^sub>n A = 0 \ 0 \ A" if "finite A" using that by (induct A) (auto simp: lcm_eq_0_iff) lemma Lcm_fin_1_iff: "Lcm\<^sub>f\<^sub>i\<^sub>n A = 1 \ (\a\A. is_unit a) \ finite A" by (induct A rule: infinite_finite_induct) simp_all end context semiring_Gcd begin lemma Gcd_fin_eq_Gcd [simp]: "Gcd\<^sub>f\<^sub>i\<^sub>n A = Gcd A" if "finite A" for A :: "'a set" using that by induct simp_all lemma Gcd_set_eq_fold [code_unfold]: "Gcd (set xs) = fold gcd xs 0" by (simp add: Gcd_fin.set_eq_fold [symmetric]) lemma Lcm_fin_eq_Lcm [simp]: "Lcm\<^sub>f\<^sub>i\<^sub>n A = Lcm A" if "finite A" for A :: "'a set" using that by induct simp_all lemma Lcm_set_eq_fold [code_unfold]: "Lcm (set xs) = fold lcm xs 1" by (simp add: Lcm_fin.set_eq_fold [symmetric]) end subsection \Coprimality\ context semiring_gcd begin lemma coprime_imp_gcd_eq_1 [simp]: "gcd a b = 1" if "coprime a b" proof - define t r s where "t = gcd a b" and "r = a div t" and "s = b div t" then have "a = t * r" and "b = t * s" by simp_all with that have "coprime (t * r) (t * s)" by simp then show ?thesis by (simp add: t_def) qed lemma gcd_eq_1_imp_coprime [dest!]: "coprime a b" if "gcd a b = 1" proof (rule coprimeI) fix c assume "c dvd a" and "c dvd b" then have "c dvd gcd a b" by (rule gcd_greatest) with that show "is_unit c" by simp qed lemma coprime_iff_gcd_eq_1 [presburger, code]: "coprime a b \ gcd a b = 1" by rule (simp_all add: gcd_eq_1_imp_coprime) lemma is_unit_gcd [simp]: "is_unit (gcd a b) \ coprime a b" by (simp add: coprime_iff_gcd_eq_1) lemma coprime_add_one_left [simp]: "coprime (a + 1) a" by (simp add: gcd_eq_1_imp_coprime ac_simps) lemma coprime_add_one_right [simp]: "coprime a (a + 1)" using coprime_add_one_left [of a] by (simp add: ac_simps) lemma coprime_mult_left_iff [simp]: "coprime (a * b) c \ coprime a c \ coprime b c" proof assume "coprime (a * b) c" with coprime_common_divisor [of "a * b" c] have *: "is_unit d" if "d dvd a * b" and "d dvd c" for d using that by blast have "coprime a c" by (rule coprimeI, rule *) simp_all moreover have "coprime b c" by (rule coprimeI, rule *) simp_all ultimately show "coprime a c \ coprime b c" .. next assume "coprime a c \ coprime b c" then have "coprime a c" "coprime b c" by simp_all show "coprime (a * b) c" proof (rule coprimeI) fix d assume "d dvd a * b" then obtain r s where d: "d = r * s" "r dvd a" "s dvd b" by (rule dvd_productE) assume "d dvd c" with d have "r * s dvd c" by simp then have "r dvd c" "s dvd c" by (auto intro: dvd_mult_left dvd_mult_right) from \coprime a c\ \r dvd a\ \r dvd c\ have "is_unit r" by (rule coprime_common_divisor) moreover from \coprime b c\ \s dvd b\ \s dvd c\ have "is_unit s" by (rule coprime_common_divisor) ultimately show "is_unit d" by (simp add: d is_unit_mult_iff) qed qed lemma coprime_mult_right_iff [simp]: "coprime c (a * b) \ coprime c a \ coprime c b" using coprime_mult_left_iff [of a b c] by (simp add: ac_simps) lemma coprime_power_left_iff [simp]: "coprime (a ^ n) b \ coprime a b \ n = 0" proof (cases "n = 0") case True then show ?thesis by simp next case False then have "n > 0" by simp then show ?thesis by (induction n rule: nat_induct_non_zero) simp_all qed lemma coprime_power_right_iff [simp]: "coprime a (b ^ n) \ coprime a b \ n = 0" using coprime_power_left_iff [of b n a] by (simp add: ac_simps) lemma prod_coprime_left: "coprime (\i\A. f i) a" if "\i. i \ A \ coprime (f i) a" using that by (induct A rule: infinite_finite_induct) simp_all lemma prod_coprime_right: "coprime a (\i\A. f i)" if "\i. i \ A \ coprime a (f i)" using that prod_coprime_left [of A f a] by (simp add: ac_simps) lemma prod_list_coprime_left: "coprime (prod_list xs) a" if "\x. x \ set xs \ coprime x a" using that by (induct xs) simp_all lemma prod_list_coprime_right: "coprime a (prod_list xs)" if "\x. x \ set xs \ coprime a x" using that prod_list_coprime_left [of xs a] by (simp add: ac_simps) lemma coprime_dvd_mult_left_iff: "a dvd b * c \ a dvd b" if "coprime a c" proof assume "a dvd b" then show "a dvd b * c" by simp next assume "a dvd b * c" show "a dvd b" proof (cases "b = 0") case True then show ?thesis by simp next case False then have unit: "is_unit (unit_factor b)" by simp from \coprime a c\ have "gcd (b * a) (b * c) * unit_factor b = b" by (subst gcd_mult_left) (simp add: ac_simps) moreover from \a dvd b * c\ have "a dvd gcd (b * a) (b * c) * unit_factor b" by (simp add: dvd_mult_unit_iff unit) ultimately show ?thesis by simp qed qed lemma coprime_dvd_mult_right_iff: "a dvd c * b \ a dvd b" if "coprime a c" using that coprime_dvd_mult_left_iff [of a c b] by (simp add: ac_simps) lemma divides_mult: "a * b dvd c" if "a dvd c" and "b dvd c" and "coprime a b" proof - from \b dvd c\ obtain b' where "c = b * b'" .. with \a dvd c\ have "a dvd b' * b" by (simp add: ac_simps) with \coprime a b\ have "a dvd b'" by (simp add: coprime_dvd_mult_left_iff) then obtain a' where "b' = a * a'" .. with \c = b * b'\ have "c = (a * b) * a'" by (simp add: ac_simps) then show ?thesis .. qed lemma div_gcd_coprime: assumes "a \ 0 \ b \ 0" shows "coprime (a div gcd a b) (b div gcd a b)" proof - let ?g = "gcd a b" let ?a' = "a div ?g" let ?b' = "b div ?g" let ?g' = "gcd ?a' ?b'" have dvdg: "?g dvd a" "?g dvd b" by simp_all have dvdg': "?g' dvd ?a'" "?g' dvd ?b'" by simp_all from dvdg dvdg' obtain ka kb ka' kb' where kab: "a = ?g * ka" "b = ?g * kb" "?a' = ?g' * ka'" "?b' = ?g' * kb'" unfolding dvd_def by blast from this [symmetric] have "?g * ?a' = (?g * ?g') * ka'" "?g * ?b' = (?g * ?g') * kb'" by (simp_all add: mult.assoc mult.left_commute [of "gcd a b"]) then have dvdgg':"?g * ?g' dvd a" "?g* ?g' dvd b" by (auto simp: dvd_mult_div_cancel [OF dvdg(1)] dvd_mult_div_cancel [OF dvdg(2)] dvd_def) have "?g \ 0" using assms by simp moreover from gcd_greatest [OF dvdgg'] have "?g * ?g' dvd ?g" . ultimately show ?thesis using dvd_times_left_cancel_iff [of "gcd a b" _ 1] by simp (simp only: coprime_iff_gcd_eq_1) qed lemma gcd_coprime: assumes c: "gcd a b \ 0" and a: "a = a' * gcd a b" and b: "b = b' * gcd a b" shows "coprime a' b'" proof - from c have "a \ 0 \ b \ 0" by simp with div_gcd_coprime have "coprime (a div gcd a b) (b div gcd a b)" . also from assms have "a div gcd a b = a'" using dvd_div_eq_mult gcd_dvd1 by blast also from assms have "b div gcd a b = b'" using dvd_div_eq_mult gcd_dvd1 by blast finally show ?thesis . qed lemma gcd_coprime_exists: assumes "gcd a b \ 0" shows "\a' b'. a = a' * gcd a b \ b = b' * gcd a b \ coprime a' b'" proof - have "coprime (a div gcd a b) (b div gcd a b)" using assms div_gcd_coprime by auto then show ?thesis by force qed lemma pow_divides_pow_iff [simp]: "a ^ n dvd b ^ n \ a dvd b" if "n > 0" proof (cases "gcd a b = 0") case True then show ?thesis by simp next case False show ?thesis proof let ?d = "gcd a b" from \n > 0\ obtain m where m: "n = Suc m" by (cases n) simp_all from False have zn: "?d ^ n \ 0" by (rule power_not_zero) from gcd_coprime_exists [OF False] obtain a' b' where ab': "a = a' * ?d" "b = b' * ?d" "coprime a' b'" by blast assume "a ^ n dvd b ^ n" then have "(a' * ?d) ^ n dvd (b' * ?d) ^ n" by (simp add: ab'(1,2)[symmetric]) then have "?d^n * a'^n dvd ?d^n * b'^n" by (simp only: power_mult_distrib ac_simps) with zn have "a' ^ n dvd b' ^ n" by simp then have "a' dvd b' ^ n" using dvd_trans[of a' "a'^n" "b'^n"] by (simp add: m) then have "a' dvd b' ^ m * b'" by (simp add: m ac_simps) moreover have "coprime a' (b' ^ n)" using \coprime a' b'\ by simp then have "a' dvd b'" using \a' dvd b' ^ n\ coprime_dvd_mult_left_iff dvd_mult by blast then have "a' * ?d dvd b' * ?d" by (rule mult_dvd_mono) simp with ab'(1,2) show "a dvd b" by simp next assume "a dvd b" with \n > 0\ show "a ^ n dvd b ^ n" by (induction rule: nat_induct_non_zero) (simp_all add: mult_dvd_mono) qed qed lemma coprime_crossproduct: fixes a b c d :: 'a assumes "coprime a d" and "coprime b c" shows "normalize a * normalize c = normalize b * normalize d \ normalize a = normalize b \ normalize c = normalize d" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by simp next assume ?lhs from \?lhs\ have "normalize a dvd normalize b * normalize d" by (auto intro: dvdI dest: sym) with \coprime a d\ have "a dvd b" by (simp add: coprime_dvd_mult_left_iff normalize_mult [symmetric]) from \?lhs\ have "normalize b dvd normalize a * normalize c" by (auto intro: dvdI dest: sym) with \coprime b c\ have "b dvd a" by (simp add: coprime_dvd_mult_left_iff normalize_mult [symmetric]) from \?lhs\ have "normalize c dvd normalize d * normalize b" by (auto intro: dvdI dest: sym simp add: mult.commute) with \coprime b c\ have "c dvd d" by (simp add: coprime_dvd_mult_left_iff coprime_commute normalize_mult [symmetric]) from \?lhs\ have "normalize d dvd normalize c * normalize a" by (auto intro: dvdI dest: sym simp add: mult.commute) with \coprime a d\ have "d dvd c" by (simp add: coprime_dvd_mult_left_iff coprime_commute normalize_mult [symmetric]) from \a dvd b\ \b dvd a\ have "normalize a = normalize b" by (rule associatedI) moreover from \c dvd d\ \d dvd c\ have "normalize c = normalize d" by (rule associatedI) ultimately show ?rhs .. qed lemma gcd_mult_left_left_cancel: "gcd (c * a) b = gcd a b" if "coprime b c" proof - have "coprime (gcd b (a * c)) c" by (rule coprimeI) (auto intro: that coprime_common_divisor) then have "gcd b (a * c) dvd a" using coprime_dvd_mult_left_iff [of "gcd b (a * c)" c a] by simp then show ?thesis by (auto intro: associated_eqI simp add: ac_simps) qed lemma gcd_mult_left_right_cancel: "gcd (a * c) b = gcd a b" if "coprime b c" using that gcd_mult_left_left_cancel [of b c a] by (simp add: ac_simps) lemma gcd_mult_right_left_cancel: "gcd a (c * b) = gcd a b" if "coprime a c" using that gcd_mult_left_left_cancel [of a c b] by (simp add: ac_simps) lemma gcd_mult_right_right_cancel: "gcd a (b * c) = gcd a b" if "coprime a c" using that gcd_mult_right_left_cancel [of a c b] by (simp add: ac_simps) lemma gcd_exp_weak: "gcd (a ^ n) (b ^ n) = normalize (gcd a b ^ n)" proof (cases "a = 0 \ b = 0 \ n = 0") case True then show ?thesis by (cases n) simp_all next case False then have "coprime (a div gcd a b) (b div gcd a b)" and "n > 0" by (auto intro: div_gcd_coprime) then have "coprime ((a div gcd a b) ^ n) ((b div gcd a b) ^ n)" by simp then have "1 = gcd ((a div gcd a b) ^ n) ((b div gcd a b) ^ n)" by simp then have "normalize (gcd a b ^ n) = normalize (gcd a b ^ n * \)" by simp also have "\ = gcd (gcd a b ^ n * (a div gcd a b) ^ n) (gcd a b ^ n * (b div gcd a b) ^ n)" by (rule gcd_mult_left [symmetric]) also have "(gcd a b) ^ n * (a div gcd a b) ^ n = a ^ n" by (simp add: ac_simps div_power dvd_power_same) also have "(gcd a b) ^ n * (b div gcd a b) ^ n = b ^ n" by (simp add: ac_simps div_power dvd_power_same) finally show ?thesis by simp qed lemma division_decomp: assumes "a dvd b * c" shows "\b' c'. a = b' * c' \ b' dvd b \ c' dvd c" proof (cases "gcd a b = 0") case True then have "a = 0 \ b = 0" by simp then have "a = 0 * c \ 0 dvd b \ c dvd c" by simp then show ?thesis by blast next case False let ?d = "gcd a b" from gcd_coprime_exists [OF False] obtain a' b' where ab': "a = a' * ?d" "b = b' * ?d" "coprime a' b'" by blast from ab'(1) have "a' dvd a" .. with assms have "a' dvd b * c" using dvd_trans [of a' a "b * c"] by simp from assms ab'(1,2) have "a' * ?d dvd (b' * ?d) * c" by simp then have "?d * a' dvd ?d * (b' * c)" by (simp add: mult_ac) with \?d \ 0\ have "a' dvd b' * c" by simp then have "a' dvd c" using \coprime a' b'\ by (simp add: coprime_dvd_mult_right_iff) with ab'(1) have "a = ?d * a' \ ?d dvd b \ a' dvd c" by (simp add: ac_simps) then show ?thesis by blast qed lemma lcm_coprime: "coprime a b \ lcm a b = normalize (a * b)" by (subst lcm_gcd) simp end context ring_gcd begin lemma coprime_minus_left_iff [simp]: "coprime (- a) b \ coprime a b" by (rule; rule coprimeI) (auto intro: coprime_common_divisor) lemma coprime_minus_right_iff [simp]: "coprime a (- b) \ coprime a b" using coprime_minus_left_iff [of b a] by (simp add: ac_simps) lemma coprime_diff_one_left [simp]: "coprime (a - 1) a" using coprime_add_one_right [of "a - 1"] by simp lemma coprime_doff_one_right [simp]: "coprime a (a - 1)" using coprime_diff_one_left [of a] by (simp add: ac_simps) end context semiring_Gcd begin lemma Lcm_coprime: assumes "finite A" and "A \ {}" and "\a b. a \ A \ b \ A \ a \ b \ coprime a b" shows "Lcm A = normalize (\A)" using assms proof (induct rule: finite_ne_induct) case singleton then show ?case by simp next case (insert a A) have "Lcm (insert a A) = lcm a (Lcm A)" by simp also from insert have "Lcm A = normalize (\A)" by blast also have "lcm a \ = lcm a (\A)" by (cases "\A = 0") (simp_all add: lcm_div_unit2) also from insert have "coprime a (\A)" by (subst coprime_commute, intro prod_coprime_left) auto with insert have "lcm a (\A) = normalize (\(insert a A))" by (simp add: lcm_coprime) finally show ?case . qed lemma Lcm_coprime': "card A \ 0 \ (\a b. a \ A \ b \ A \ a \ b \ coprime a b) \ Lcm A = normalize (\A)" by (rule Lcm_coprime) (simp_all add: card_eq_0_iff) end subsection \GCD and LCM for multiplicative normalisation functions\ class semiring_gcd_mult_normalize = semiring_gcd + normalization_semidom_multiplicative begin lemma mult_gcd_left: "c * gcd a b = unit_factor c * gcd (c * a) (c * b)" by (simp add: gcd_mult_left normalize_mult mult.assoc [symmetric]) lemma mult_gcd_right: "gcd a b * c = gcd (a * c) (b * c) * unit_factor c" using mult_gcd_left [of c a b] by (simp add: ac_simps) lemma gcd_mult_distrib': "normalize c * gcd a b = gcd (c * a) (c * b)" by (subst gcd_mult_left) (simp_all add: normalize_mult) lemma gcd_mult_distrib: "k * gcd a b = gcd (k * a) (k * b) * unit_factor k" proof- have "normalize k * gcd a b = gcd (k * a) (k * b)" by (simp add: gcd_mult_distrib') then have "normalize k * gcd a b * unit_factor k = gcd (k * a) (k * b) * unit_factor k" by simp then have "normalize k * unit_factor k * gcd a b = gcd (k * a) (k * b) * unit_factor k" by (simp only: ac_simps) then show ?thesis by simp qed lemma gcd_mult_lcm [simp]: "gcd a b * lcm a b = normalize a * normalize b" by (simp add: lcm_gcd normalize_mult dvd_normalize_div) lemma lcm_mult_gcd [simp]: "lcm a b * gcd a b = normalize a * normalize b" using gcd_mult_lcm [of a b] by (simp add: ac_simps) lemma mult_lcm_left: "c * lcm a b = unit_factor c * lcm (c * a) (c * b)" by (simp add: lcm_mult_left mult.assoc [symmetric] normalize_mult) lemma mult_lcm_right: "lcm a b * c = lcm (a * c) (b * c) * unit_factor c" using mult_lcm_left [of c a b] by (simp add: ac_simps) lemma lcm_gcd_prod: "lcm a b * gcd a b = normalize (a * b)" by (simp add: lcm_gcd dvd_normalize_div normalize_mult) lemma lcm_mult_distrib': "normalize c * lcm a b = lcm (c * a) (c * b)" by (subst lcm_mult_left) (simp add: normalize_mult) lemma lcm_mult_distrib: "k * lcm a b = lcm (k * a) (k * b) * unit_factor k" proof- have "normalize k * lcm a b = lcm (k * a) (k * b)" by (simp add: lcm_mult_distrib') then have "normalize k * lcm a b * unit_factor k = lcm (k * a) (k * b) * unit_factor k" by simp then have "normalize k * unit_factor k * lcm a b = lcm (k * a) (k * b) * unit_factor k" by (simp only: ac_simps) then show ?thesis by simp qed lemma coprime_crossproduct': fixes a b c d assumes "b \ 0" assumes unit_factors: "unit_factor b = unit_factor d" assumes coprime: "coprime a b" "coprime c d" shows "a * d = b * c \ a = c \ b = d" proof safe assume eq: "a * d = b * c" hence "normalize a * normalize d = normalize c * normalize b" by (simp only: normalize_mult [symmetric] mult_ac) with coprime have "normalize b = normalize d" by (subst (asm) coprime_crossproduct) simp_all from this and unit_factors show "b = d" by (rule normalize_unit_factor_eqI) from eq have "a * d = c * d" by (simp only: \b = d\ mult_ac) with \b \ 0\ \b = d\ show "a = c" by simp qed (simp_all add: mult_ac) lemma gcd_exp [simp]: "gcd (a ^ n) (b ^ n) = gcd a b ^ n" using gcd_exp_weak[of a n b] by (simp add: normalize_power) end subsection \GCD and LCM on \<^typ>\nat\ and \<^typ>\int\\ instantiation nat :: gcd begin fun gcd_nat :: "nat \ nat \ nat" where "gcd_nat x y = (if y = 0 then x else gcd y (x mod y))" definition lcm_nat :: "nat \ nat \ nat" where "lcm_nat x y = x * y div (gcd x y)" instance .. end instantiation int :: gcd begin definition gcd_int :: "int \ int \ int" where "gcd_int x y = int (gcd (nat \x\) (nat \y\))" definition lcm_int :: "int \ int \ int" where "lcm_int x y = int (lcm (nat \x\) (nat \y\))" instance .. end lemma gcd_int_int_eq [simp]: "gcd (int m) (int n) = int (gcd m n)" by (simp add: gcd_int_def) lemma gcd_nat_abs_left_eq [simp]: "gcd (nat \k\) n = nat (gcd k (int n))" by (simp add: gcd_int_def) lemma gcd_nat_abs_right_eq [simp]: "gcd n (nat \k\) = nat (gcd (int n) k)" by (simp add: gcd_int_def) lemma abs_gcd_int [simp]: "\gcd x y\ = gcd x y" for x y :: int by (simp only: gcd_int_def) lemma gcd_abs1_int [simp]: "gcd \x\ y = gcd x y" for x y :: int by (simp only: gcd_int_def) simp lemma gcd_abs2_int [simp]: "gcd x \y\ = gcd x y" for x y :: int by (simp only: gcd_int_def) simp lemma lcm_int_int_eq [simp]: "lcm (int m) (int n) = int (lcm m n)" by (simp add: lcm_int_def) lemma lcm_nat_abs_left_eq [simp]: "lcm (nat \k\) n = nat (lcm k (int n))" by (simp add: lcm_int_def) lemma lcm_nat_abs_right_eq [simp]: "lcm n (nat \k\) = nat (lcm (int n) k)" by (simp add: lcm_int_def) lemma lcm_abs1_int [simp]: "lcm \x\ y = lcm x y" for x y :: int by (simp only: lcm_int_def) simp lemma lcm_abs2_int [simp]: "lcm x \y\ = lcm x y" for x y :: int by (simp only: lcm_int_def) simp lemma abs_lcm_int [simp]: "\lcm i j\ = lcm i j" for i j :: int by (simp only: lcm_int_def) lemma gcd_nat_induct [case_names base step]: fixes m n :: nat assumes "\m. P m 0" and "\m n. 0 < n \ P n (m mod n) \ P m n" shows "P m n" proof (induction m n rule: gcd_nat.induct) case (1 x y) then show ?case using assms neq0_conv by blast qed lemma gcd_neg1_int [simp]: "gcd (- x) y = gcd x y" for x y :: int by (simp only: gcd_int_def) simp lemma gcd_neg2_int [simp]: "gcd x (- y) = gcd x y" for x y :: int by (simp only: gcd_int_def) simp lemma gcd_cases_int: fixes x y :: int assumes "x \ 0 \ y \ 0 \ P (gcd x y)" and "x \ 0 \ y \ 0 \ P (gcd x (- y))" and "x \ 0 \ y \ 0 \ P (gcd (- x) y)" and "x \ 0 \ y \ 0 \ P (gcd (- x) (- y))" shows "P (gcd x y)" using assms by auto arith lemma gcd_ge_0_int [simp]: "gcd (x::int) y >= 0" for x y :: int by (simp add: gcd_int_def) lemma lcm_neg1_int: "lcm (- x) y = lcm x y" for x y :: int by (simp only: lcm_int_def) simp lemma lcm_neg2_int: "lcm x (- y) = lcm x y" for x y :: int by (simp only: lcm_int_def) simp lemma lcm_cases_int: fixes x y :: int assumes "x \ 0 \ y \ 0 \ P (lcm x y)" and "x \ 0 \ y \ 0 \ P (lcm x (- y))" and "x \ 0 \ y \ 0 \ P (lcm (- x) y)" and "x \ 0 \ y \ 0 \ P (lcm (- x) (- y))" shows "P (lcm x y)" using assms by (auto simp: lcm_neg1_int lcm_neg2_int) arith lemma lcm_ge_0_int [simp]: "lcm x y \ 0" for x y :: int by (simp only: lcm_int_def) lemma gcd_0_nat: "gcd x 0 = x" for x :: nat by simp lemma gcd_0_int [simp]: "gcd x 0 = \x\" for x :: int by (auto simp: gcd_int_def) lemma gcd_0_left_nat: "gcd 0 x = x" for x :: nat by simp lemma gcd_0_left_int [simp]: "gcd 0 x = \x\" for x :: int by (auto simp: gcd_int_def) lemma gcd_red_nat: "gcd x y = gcd y (x mod y)" for x y :: nat by (cases "y = 0") auto text \Weaker, but useful for the simplifier.\ lemma gcd_non_0_nat: "y \ 0 \ gcd x y = gcd y (x mod y)" for x y :: nat by simp lemma gcd_1_nat [simp]: "gcd m 1 = 1" for m :: nat by simp lemma gcd_Suc_0 [simp]: "gcd m (Suc 0) = Suc 0" for m :: nat by simp lemma gcd_1_int [simp]: "gcd m 1 = 1" for m :: int by (simp add: gcd_int_def) lemma gcd_idem_nat: "gcd x x = x" for x :: nat by simp lemma gcd_idem_int: "gcd x x = \x\" for x :: int by (auto simp: gcd_int_def) declare gcd_nat.simps [simp del] text \ \<^medskip> \<^term>\gcd m n\ divides \m\ and \n\. The conjunctions don't seem provable separately. \ instance nat :: semiring_gcd proof fix m n :: nat show "gcd m n dvd m" and "gcd m n dvd n" proof (induct m n rule: gcd_nat_induct) case (step m n) then have "gcd n (m mod n) dvd m" by (metis dvd_mod_imp_dvd) with step show "gcd m n dvd m" by (simp add: gcd_non_0_nat) qed (simp_all add: gcd_0_nat gcd_non_0_nat) next fix m n k :: nat assume "k dvd m" and "k dvd n" then show "k dvd gcd m n" by (induct m n rule: gcd_nat_induct) (simp_all add: gcd_non_0_nat dvd_mod gcd_0_nat) qed (simp_all add: lcm_nat_def) instance int :: ring_gcd proof fix k l r :: int show [simp]: "gcd k l dvd k" "gcd k l dvd l" using gcd_dvd1 [of "nat \k\" "nat \l\"] gcd_dvd2 [of "nat \k\" "nat \l\"] by simp_all show "lcm k l = normalize (k * l div gcd k l)" using lcm_gcd [of "nat \k\" "nat \l\"] by (simp add: nat_eq_iff of_nat_div abs_mult abs_div) assume "r dvd k" "r dvd l" then show "r dvd gcd k l" using gcd_greatest [of "nat \r\" "nat \k\" "nat \l\"] by simp qed simp lemma gcd_le1_nat [simp]: "a \ 0 \ gcd a b \ a" for a b :: nat by (rule dvd_imp_le) auto lemma gcd_le2_nat [simp]: "b \ 0 \ gcd a b \ b" for a b :: nat by (rule dvd_imp_le) auto lemma gcd_le1_int [simp]: "a > 0 \ gcd a b \ a" for a b :: int by (rule zdvd_imp_le) auto lemma gcd_le2_int [simp]: "b > 0 \ gcd a b \ b" for a b :: int by (rule zdvd_imp_le) auto lemma gcd_pos_nat [simp]: "gcd m n > 0 \ m \ 0 \ n \ 0" for m n :: nat using gcd_eq_0_iff [of m n] by arith lemma gcd_pos_int [simp]: "gcd m n > 0 \ m \ 0 \ n \ 0" for m n :: int using gcd_eq_0_iff [of m n] gcd_ge_0_int [of m n] by arith lemma gcd_unique_nat: "d dvd a \ d dvd b \ (\e. e dvd a \ e dvd b \ e dvd d) \ d = gcd a b" for d a :: nat using gcd_unique by fastforce lemma gcd_unique_int: "d \ 0 \ d dvd a \ d dvd b \ (\e. e dvd a \ e dvd b \ e dvd d) \ d = gcd a b" for d a :: int using zdvd_antisym_nonneg by auto interpretation gcd_nat: semilattice_neutr_order gcd "0::nat" Rings.dvd "\m n. m dvd n \ m \ n" by standard (auto simp: gcd_unique_nat [symmetric] intro: dvd_antisym dvd_trans) lemma gcd_proj1_if_dvd_int [simp]: "x dvd y \ gcd x y = \x\" for x y :: int by (metis abs_dvd_iff gcd_0_left_int gcd_unique_int) lemma gcd_proj2_if_dvd_int [simp]: "y dvd x \ gcd x y = \y\" for x y :: int by (metis gcd_proj1_if_dvd_int gcd.commute) text \\<^medskip> Multiplication laws.\ lemma gcd_mult_distrib_nat: "k * gcd m n = gcd (k * m) (k * n)" for k m n :: nat \ \@{cite \page 27\ davenport92}\ by (simp add: gcd_mult_left) lemma gcd_mult_distrib_int: "\k\ * gcd m n = gcd (k * m) (k * n)" for k m n :: int by (simp add: gcd_mult_left abs_mult) text \\medskip Addition laws.\ (* TODO: add the other variations? *) lemma gcd_diff1_nat: "m \ n \ gcd (m - n) n = gcd m n" for m n :: nat by (subst gcd_add1 [symmetric]) auto lemma gcd_diff2_nat: "n \ m \ gcd (n - m) n = gcd m n" for m n :: nat by (metis gcd.commute gcd_add2 gcd_diff1_nat le_add_diff_inverse2) lemma gcd_non_0_int: fixes x y :: int assumes "y > 0" shows "gcd x y = gcd y (x mod y)" proof (cases "x mod y = 0") case False then have neg: "x mod y = y - (- x) mod y" by (simp add: zmod_zminus1_eq_if) have xy: "0 \ x mod y" by (simp add: assms) show ?thesis proof (cases "x < 0") case True have "nat (- x mod y) \ nat y" by (simp add: assms dual_order.order_iff_strict) moreover have "gcd (nat (- x)) (nat y) = gcd (nat (- x mod y)) (nat y)" using True assms gcd_non_0_nat nat_mod_distrib by auto ultimately have "gcd (nat (- x)) (nat y) = gcd (nat y) (nat (x mod y))" using assms by (simp add: neg nat_diff_distrib') (metis gcd.commute gcd_diff2_nat) with assms \0 \ x mod y\ show ?thesis by (simp add: True dual_order.order_iff_strict gcd_int_def) next case False with assms xy have "gcd (nat x) (nat y) = gcd (nat y) (nat x mod nat y)" using gcd_red_nat by blast with False assms show ?thesis by (simp add: gcd_int_def nat_mod_distrib) qed qed (use assms in auto) lemma gcd_red_int: "gcd x y = gcd y (x mod y)" for x y :: int proof (cases y "0::int" rule: linorder_cases) case less with gcd_non_0_int [of "- y" "- x"] show ?thesis by auto next case greater with gcd_non_0_int [of y x] show ?thesis by auto qed auto (* TODO: differences, and all variations of addition rules as simplification rules for nat and int *) (* TODO: add the three variations of these, and for ints? *) lemma finite_divisors_nat [simp]: (* FIXME move *) fixes m :: nat assumes "m > 0" shows "finite {d. d dvd m}" proof- from assms have "{d. d dvd m} \ {d. d \ m}" by (auto dest: dvd_imp_le) then show ?thesis using finite_Collect_le_nat by (rule finite_subset) qed lemma finite_divisors_int [simp]: fixes i :: int assumes "i \ 0" shows "finite {d. d dvd i}" proof - have "{d. \d\ \ \i\} = {- \i\..\i\}" by (auto simp: abs_if) then have "finite {d. \d\ \ \i\}" by simp from finite_subset [OF _ this] show ?thesis using assms by (simp add: dvd_imp_le_int subset_iff) qed lemma Max_divisors_self_nat [simp]: "n \ 0 \ Max {d::nat. d dvd n} = n" by (fastforce intro: antisym Max_le_iff[THEN iffD2] simp: dvd_imp_le) lemma Max_divisors_self_int [simp]: assumes "n \ 0" shows "Max {d::int. d dvd n} = \n\" proof (rule antisym) show "Max {d. d dvd n} \ \n\" using assms by (auto intro: abs_le_D1 dvd_imp_le_int intro!: Max_le_iff [THEN iffD2]) qed (simp add: assms) lemma gcd_is_Max_divisors_nat: fixes m n :: nat assumes "n > 0" shows "gcd m n = Max {d. d dvd m \ d dvd n}" proof (rule Max_eqI[THEN sym], simp_all) show "finite {d. d dvd m \ d dvd n}" by (simp add: \n > 0\) show "\y. y dvd m \ y dvd n \ y \ gcd m n" by (simp add: \n > 0\ dvd_imp_le) qed lemma gcd_is_Max_divisors_int: fixes m n :: int assumes "n \ 0" shows "gcd m n = Max {d. d dvd m \ d dvd n}" proof (rule Max_eqI[THEN sym], simp_all) show "finite {d. d dvd m \ d dvd n}" by (simp add: \n \ 0\) show "\y. y dvd m \ y dvd n \ y \ gcd m n" by (simp add: \n \ 0\ zdvd_imp_le) qed lemma gcd_code_int [code]: "gcd k l = \if l = 0 then k else gcd l (\k\ mod \l\)\" for k l :: int using gcd_red_int [of "\k\" "\l\"] by simp lemma coprime_Suc_left_nat [simp]: "coprime (Suc n) n" using coprime_add_one_left [of n] by simp lemma coprime_Suc_right_nat [simp]: "coprime n (Suc n)" using coprime_Suc_left_nat [of n] by (simp add: ac_simps) lemma coprime_diff_one_left_nat [simp]: "coprime (n - 1) n" if "n > 0" for n :: nat using that coprime_Suc_right_nat [of "n - 1"] by simp lemma coprime_diff_one_right_nat [simp]: "coprime n (n - 1)" if "n > 0" for n :: nat using that coprime_diff_one_left_nat [of n] by (simp add: ac_simps) lemma coprime_crossproduct_nat: fixes a b c d :: nat assumes "coprime a d" and "coprime b c" shows "a * c = b * d \ a = b \ c = d" using assms coprime_crossproduct [of a d b c] by simp lemma coprime_crossproduct_int: fixes a b c d :: int assumes "coprime a d" and "coprime b c" shows "\a\ * \c\ = \b\ * \d\ \ \a\ = \b\ \ \c\ = \d\" using assms coprime_crossproduct [of a d b c] by simp subsection \Bezout's theorem\ text \ Function \bezw\ returns a pair of witnesses to Bezout's theorem -- see the theorems that follow the definition. \ fun bezw :: "nat \ nat \ int * int" where "bezw x y = (if y = 0 then (1, 0) else (snd (bezw y (x mod y)), fst (bezw y (x mod y)) - snd (bezw y (x mod y)) * int(x div y)))" lemma bezw_0 [simp]: "bezw x 0 = (1, 0)" by simp lemma bezw_non_0: "y > 0 \ bezw x y = (snd (bezw y (x mod y)), fst (bezw y (x mod y)) - snd (bezw y (x mod y)) * int(x div y))" by simp declare bezw.simps [simp del] lemma bezw_aux: "int (gcd x y) = fst (bezw x y) * int x + snd (bezw x y) * int y" proof (induct x y rule: gcd_nat_induct) case (step m n) then have "fst (bezw m n) * int m + snd (bezw m n) * int n - int (gcd m n) = int m * snd (bezw n (m mod n)) - (int (m mod n) * snd (bezw n (m mod n)) + int n * (int (m div n) * snd (bezw n (m mod n))))" by (simp add: bezw_non_0 gcd_non_0_nat field_simps) also have "\ = int m * snd (bezw n (m mod n)) - (int (m mod n) + int (n * (m div n))) * snd (bezw n (m mod n))" by (simp add: distrib_right) also have "\ = 0" by (metis cancel_comm_monoid_add_class.diff_cancel mod_mult_div_eq of_nat_add) finally show ?case by simp qed auto lemma bezout_int: "\u v. u * x + v * y = gcd x y" for x y :: int proof - have aux: "x \ 0 \ y \ 0 \ \u v. u * x + v * y = gcd x y" for x y :: int apply (rule_tac x = "fst (bezw (nat x) (nat y))" in exI) apply (rule_tac x = "snd (bezw (nat x) (nat y))" in exI) by (simp add: bezw_aux gcd_int_def) consider "x \ 0" "y \ 0" | "x \ 0" "y \ 0" | "x \ 0" "y \ 0" | "x \ 0" "y \ 0" using linear by blast then show ?thesis proof cases case 1 then show ?thesis by (rule aux) next case 2 then show ?thesis using aux [of x "-y"] by (metis gcd_neg2_int mult.commute mult_minus_right neg_0_le_iff_le) next case 3 then show ?thesis using aux [of "-x" y] by (metis gcd.commute gcd_neg2_int mult.commute mult_minus_right neg_0_le_iff_le) next case 4 then show ?thesis using aux [of "-x" "-y"] by (metis diff_0 diff_ge_0_iff_ge gcd_neg1_int gcd_neg2_int mult.commute mult_minus_right) qed qed text \Versions of Bezout for \nat\, by Amine Chaieb.\ lemma Euclid_induct [case_names swap zero add]: fixes P :: "nat \ nat \ bool" assumes c: "\a b. P a b \ P b a" and z: "\a. P a 0" and add: "\a b. P a b \ P a (a + b)" shows "P a b" proof (induct "a + b" arbitrary: a b rule: less_induct) case less consider (eq) "a = b" | (lt) "a < b" "a + b - a < a + b" | "b = 0" | "b + a - b < a + b" by arith show ?case proof (cases a b rule: linorder_cases) case equal with add [rule_format, OF z [rule_format, of a]] show ?thesis by simp next case lt: less then consider "a = 0" | "a + b - a < a + b" by arith then show ?thesis proof cases case 1 with z c show ?thesis by blast next case 2 also have *: "a + b - a = a + (b - a)" using lt by arith finally have "a + (b - a) < a + b" . then have "P a (a + (b - a))" by (rule add [rule_format, OF less]) then show ?thesis by (simp add: *[symmetric]) qed next case gt: greater then consider "b = 0" | "b + a - b < a + b" by arith then show ?thesis proof cases case 1 with z c show ?thesis by blast next case 2 also have *: "b + a - b = b + (a - b)" using gt by arith finally have "b + (a - b) < a + b" . then have "P b (b + (a - b))" by (rule add [rule_format, OF less]) then have "P b a" by (simp add: *[symmetric]) with c show ?thesis by blast qed qed qed lemma bezout_lemma_nat: fixes d::nat shows "\d dvd a; d dvd b; a * x = b * y + d \ b * x = a * y + d\ \ \x y. d dvd a \ d dvd a + b \ (a * x = (a + b) * y + d \ (a + b) * x = a * y + d)" apply auto apply (metis add_mult_distrib2 left_add_mult_distrib) apply (rule_tac x=x in exI) by (metis add_mult_distrib2 mult.commute add.assoc) lemma bezout_add_nat: "\(d::nat) x y. d dvd a \ d dvd b \ (a * x = b * y + d \ b * x = a * y + d)" proof (induct a b rule: Euclid_induct) case (swap a b) then show ?case by blast next case (zero a) then show ?case by fastforce next case (add a b) then show ?case by (meson bezout_lemma_nat) qed lemma bezout1_nat: "\(d::nat) x y. d dvd a \ d dvd b \ (a * x - b * y = d \ b * x - a * y = d)" using bezout_add_nat[of a b] by (metis add_diff_cancel_left') lemma bezout_add_strong_nat: fixes a b :: nat assumes a: "a \ 0" shows "\d x y. d dvd a \ d dvd b \ a * x = b * y + d" proof - consider d x y where "d dvd a" "d dvd b" "a * x = b * y + d" | d x y where "d dvd a" "d dvd b" "b * x = a * y + d" using bezout_add_nat [of a b] by blast then show ?thesis proof cases case 1 then show ?thesis by blast next case H: 2 show ?thesis proof (cases "b = 0") case True with H show ?thesis by simp next case False then have bp: "b > 0" by simp with dvd_imp_le [OF H(2)] consider "d = b" | "d < b" by atomize_elim auto then show ?thesis proof cases case 1 with a H show ?thesis by (metis Suc_pred add.commute mult.commute mult_Suc_right neq0_conv) next case 2 show ?thesis proof (cases "x = 0") case True with a H show ?thesis by simp next case x0: False then have xp: "x > 0" by simp from \d < b\ have "d \ b - 1" by simp then have "d * b \ b * (b - 1)" by simp with xp mult_mono[of "1" "x" "d * b" "b * (b - 1)"] have dble: "d * b \ x * b * (b - 1)" using bp by simp from H(3) have "d + (b - 1) * (b * x) = d + (b - 1) * (a * y + d)" by simp then have "d + (b - 1) * a * y + (b - 1) * d = d + (b - 1) * b * x" by (simp only: mult.assoc distrib_left) then have "a * ((b - 1) * y) + d * (b - 1 + 1) = d + x * b * (b - 1)" by algebra then have "a * ((b - 1) * y) = d + x * b * (b - 1) - d * b" using bp by simp then have "a * ((b - 1) * y) = d + (x * b * (b - 1) - d * b)" by (simp only: diff_add_assoc[OF dble, of d, symmetric]) then have "a * ((b - 1) * y) = b * (x * (b - 1) - d) + d" by (simp only: diff_mult_distrib2 ac_simps) with H(1,2) show ?thesis by blast qed qed qed qed qed lemma bezout_nat: fixes a :: nat assumes a: "a \ 0" shows "\x y. a * x = b * y + gcd a b" proof - obtain d x y where d: "d dvd a" "d dvd b" and eq: "a * x = b * y + d" using bezout_add_strong_nat [OF a, of b] by blast from d have "d dvd gcd a b" by simp then obtain k where k: "gcd a b = d * k" unfolding dvd_def by blast from eq have "a * x * k = (b * y + d) * k" by auto then have "a * (x * k) = b * (y * k) + gcd a b" by (algebra add: k) then show ?thesis by blast qed subsection \LCM properties on \<^typ>\nat\ and \<^typ>\int\\ lemma lcm_altdef_int [code]: "lcm a b = \a\ * \b\ div gcd a b" for a b :: int by (simp add: abs_mult lcm_gcd abs_div) lemma prod_gcd_lcm_nat: "m * n = gcd m n * lcm m n" for m n :: nat by (simp add: lcm_gcd) lemma prod_gcd_lcm_int: "\m\ * \n\ = gcd m n * lcm m n" for m n :: int by (simp add: lcm_gcd abs_div abs_mult) lemma lcm_pos_nat: "m > 0 \ n > 0 \ lcm m n > 0" for m n :: nat using lcm_eq_0_iff [of m n] by auto lemma lcm_pos_int: "m \ 0 \ n \ 0 \ lcm m n > 0" for m n :: int by (simp add: less_le lcm_eq_0_iff) lemma dvd_pos_nat: "n > 0 \ m dvd n \ m > 0" (* FIXME move *) for m n :: nat by auto lemma lcm_unique_nat: "a dvd d \ b dvd d \ (\e. a dvd e \ b dvd e \ d dvd e) \ d = lcm a b" for a b d :: nat by (auto intro: dvd_antisym lcm_least) lemma lcm_unique_int: "d \ 0 \ a dvd d \ b dvd d \ (\e. a dvd e \ b dvd e \ d dvd e) \ d = lcm a b" for a b d :: int using lcm_least zdvd_antisym_nonneg by auto lemma lcm_proj2_if_dvd_nat [simp]: "x dvd y \ lcm x y = y" for x y :: nat by (simp add: lcm_proj2_if_dvd) lemma lcm_proj2_if_dvd_int [simp]: "x dvd y \ lcm x y = \y\" for x y :: int by (simp add: lcm_proj2_if_dvd) lemma lcm_proj1_if_dvd_nat [simp]: "x dvd y \ lcm y x = y" for x y :: nat by (subst lcm.commute) (erule lcm_proj2_if_dvd_nat) lemma lcm_proj1_if_dvd_int [simp]: "x dvd y \ lcm y x = \y\" for x y :: int by (subst lcm.commute) (erule lcm_proj2_if_dvd_int) lemma lcm_proj1_iff_nat [simp]: "lcm m n = m \ n dvd m" for m n :: nat by (metis lcm_proj1_if_dvd_nat lcm_unique_nat) lemma lcm_proj2_iff_nat [simp]: "lcm m n = n \ m dvd n" for m n :: nat by (metis lcm_proj2_if_dvd_nat lcm_unique_nat) lemma lcm_proj1_iff_int [simp]: "lcm m n = \m\ \ n dvd m" for m n :: int by (metis dvd_abs_iff lcm_proj1_if_dvd_int lcm_unique_int) lemma lcm_proj2_iff_int [simp]: "lcm m n = \n\ \ m dvd n" for m n :: int by (metis dvd_abs_iff lcm_proj2_if_dvd_int lcm_unique_int) lemma lcm_1_iff_nat [simp]: "lcm m n = Suc 0 \ m = Suc 0 \ n = Suc 0" for m n :: nat using lcm_eq_1_iff [of m n] by simp lemma lcm_1_iff_int [simp]: "lcm m n = 1 \ (m = 1 \ m = -1) \ (n = 1 \ n = -1)" for m n :: int by auto subsection \The complete divisibility lattice on \<^typ>\nat\ and \<^typ>\int\\ text \ Lifting \gcd\ and \lcm\ to sets (\Gcd\ / \Lcm\). \Gcd\ is defined via \Lcm\ to facilitate the proof that we have a complete lattice. \ instantiation nat :: semiring_Gcd begin interpretation semilattice_neutr_set lcm "1::nat" by standard simp_all definition "Lcm M = (if finite M then F M else 0)" for M :: "nat set" lemma Lcm_nat_empty: "Lcm {} = (1::nat)" by (simp add: Lcm_nat_def del: One_nat_def) lemma Lcm_nat_insert: "Lcm (insert n M) = lcm n (Lcm M)" for n :: nat by (cases "finite M") (auto simp: Lcm_nat_def simp del: One_nat_def) lemma Lcm_nat_infinite: "infinite M \ Lcm M = 0" for M :: "nat set" by (simp add: Lcm_nat_def) lemma dvd_Lcm_nat [simp]: fixes M :: "nat set" assumes "m \ M" shows "m dvd Lcm M" proof - from assms have "insert m M = M" by auto moreover have "m dvd Lcm (insert m M)" by (simp add: Lcm_nat_insert) ultimately show ?thesis by simp qed lemma Lcm_dvd_nat [simp]: fixes M :: "nat set" assumes "\m\M. m dvd n" shows "Lcm M dvd n" proof (cases "n > 0") case False then show ?thesis by simp next case True then have "finite {d. d dvd n}" by (rule finite_divisors_nat) moreover have "M \ {d. d dvd n}" using assms by fast ultimately have "finite M" by (rule rev_finite_subset) then show ?thesis using assms by (induct M) (simp_all add: Lcm_nat_empty Lcm_nat_insert) qed definition "Gcd M = Lcm {d. \m\M. d dvd m}" for M :: "nat set" instance proof fix N :: "nat set" fix n :: nat show "Gcd N dvd n" if "n \ N" using that by (induct N rule: infinite_finite_induct) (auto simp: Gcd_nat_def) show "n dvd Gcd N" if "\m. m \ N \ n dvd m" using that by (induct N rule: infinite_finite_induct) (auto simp: Gcd_nat_def) show "n dvd Lcm N" if "n \ N" using that by (induct N rule: infinite_finite_induct) auto show "Lcm N dvd n" if "\m. m \ N \ m dvd n" using that by (induct N rule: infinite_finite_induct) auto show "normalize (Gcd N) = Gcd N" and "normalize (Lcm N) = Lcm N" by simp_all qed end lemma Gcd_nat_eq_one: "1 \ N \ Gcd N = 1" for N :: "nat set" by (rule Gcd_eq_1_I) auto instance nat :: semiring_gcd_mult_normalize by intro_classes (auto simp: unit_factor_nat_def) text \Alternative characterizations of Gcd:\ lemma Gcd_eq_Max: fixes M :: "nat set" assumes "finite (M::nat set)" and "M \ {}" and "0 \ M" shows "Gcd M = Max (\m\M. {d. d dvd m})" proof (rule antisym) from assms obtain m where "m \ M" and "m > 0" by auto from \m > 0\ have "finite {d. d dvd m}" by (blast intro: finite_divisors_nat) with \m \ M\ have fin: "finite (\m\M. {d. d dvd m})" by blast from fin show "Gcd M \ Max (\m\M. {d. d dvd m})" by (auto intro: Max_ge Gcd_dvd) from fin show "Max (\m\M. {d. d dvd m}) \ Gcd M" proof (rule Max.boundedI, simp_all) show "(\m\M. {d. d dvd m}) \ {}" by auto show "\a. \x\M. a dvd x \ a \ Gcd M" by (meson Gcd_dvd Gcd_greatest \0 < m\ \m \ M\ dvd_imp_le dvd_pos_nat) qed qed lemma Gcd_remove0_nat: "finite M \ Gcd M = Gcd (M - {0})" for M :: "nat set" proof (induct pred: finite) case (insert x M) then show ?case by (simp add: insert_Diff_if) qed auto lemma Lcm_in_lcm_closed_set_nat: fixes M :: "nat set" assumes "finite M" "M \ {}" "\m n. \m \ M; n \ M\ \ lcm m n \ M" shows "Lcm M \ M" using assms proof (induction M rule: finite_linorder_min_induct) case (insert x M) then have "\m n. m \ M \ n \ M \ lcm m n \ M" by (metis dvd_lcm1 gr0I insert_iff lcm_pos_nat nat_dvd_not_less) with insert show ?case by simp (metis Lcm_nat_empty One_nat_def dvd_1_left dvd_lcm2) qed auto lemma Lcm_eq_Max_nat: fixes M :: "nat set" assumes M: "finite M" "M \ {}" "0 \ M" and lcm: "\m n. \m \ M; n \ M\ \ lcm m n \ M" shows "Lcm M = Max M" proof (rule antisym) show "Lcm M \ Max M" by (simp add: Lcm_in_lcm_closed_set_nat \finite M\ \M \ {}\ lcm) show "Max M \ Lcm M" by (meson Lcm_0_iff Max_in M dvd_Lcm dvd_imp_le le_0_eq not_le) qed lemma mult_inj_if_coprime_nat: "inj_on f A \ inj_on g B \ (\a b. \a\A; b\B\ \ coprime (f a) (g b)) \ inj_on (\(a, b). f a * g b) (A \ B)" for f :: "'a \ nat" and g :: "'b \ nat" by (auto simp: inj_on_def coprime_crossproduct_nat simp del: One_nat_def) subsubsection \Setwise GCD and LCM for integers\ instantiation int :: Gcd begin definition Gcd_int :: "int set \ int" where "Gcd K = int (GCD k\K. (nat \ abs) k)" definition Lcm_int :: "int set \ int" where "Lcm K = int (LCM k\K. (nat \ abs) k)" instance .. end lemma Gcd_int_eq [simp]: "(GCD n\N. int n) = int (Gcd N)" by (simp add: Gcd_int_def image_image) lemma Gcd_nat_abs_eq [simp]: "(GCD k\K. nat \k\) = nat (Gcd K)" by (simp add: Gcd_int_def) lemma abs_Gcd_eq [simp]: "\Gcd K\ = Gcd K" for K :: "int set" by (simp only: Gcd_int_def) lemma Gcd_int_greater_eq_0 [simp]: "Gcd K \ 0" for K :: "int set" using abs_ge_zero [of "Gcd K"] by simp lemma Gcd_abs_eq [simp]: "(GCD k\K. \k\) = Gcd K" for K :: "int set" by (simp only: Gcd_int_def image_image) simp lemma Lcm_int_eq [simp]: "(LCM n\N. int n) = int (Lcm N)" by (simp add: Lcm_int_def image_image) lemma Lcm_nat_abs_eq [simp]: "(LCM k\K. nat \k\) = nat (Lcm K)" by (simp add: Lcm_int_def) lemma abs_Lcm_eq [simp]: "\Lcm K\ = Lcm K" for K :: "int set" by (simp only: Lcm_int_def) lemma Lcm_int_greater_eq_0 [simp]: "Lcm K \ 0" for K :: "int set" using abs_ge_zero [of "Lcm K"] by simp lemma Lcm_abs_eq [simp]: "(LCM k\K. \k\) = Lcm K" for K :: "int set" by (simp only: Lcm_int_def image_image) simp instance int :: semiring_Gcd proof fix K :: "int set" and k :: int show "Gcd K dvd k" and "k dvd Lcm K" if "k \ K" using that Gcd_dvd [of "nat \k\" "(nat \ abs) ` K"] dvd_Lcm [of "nat \k\" "(nat \ abs) ` K"] by (simp_all add: comp_def) show "k dvd Gcd K" if "\l. l \ K \ k dvd l" proof - have "nat \k\ dvd (GCD k\K. nat \k\)" by (rule Gcd_greatest) (use that in auto) then show ?thesis by simp qed show "Lcm K dvd k" if "\l. l \ K \ l dvd k" proof - have "(LCM k\K. nat \k\) dvd nat \k\" by (rule Lcm_least) (use that in auto) then show ?thesis by simp qed qed (simp_all add: sgn_mult) instance int :: semiring_gcd_mult_normalize by intro_classes (auto simp: sgn_mult) subsection \GCD and LCM on \<^typ>\integer\\ instantiation integer :: gcd begin context includes integer.lifting begin lift_definition gcd_integer :: "integer \ integer \ integer" is gcd . lift_definition lcm_integer :: "integer \ integer \ integer" is lcm . end instance .. end lifting_update integer.lifting lifting_forget integer.lifting context includes integer.lifting begin lemma gcd_code_integer [code]: "gcd k l = \if l = (0::integer) then k else gcd l (\k\ mod \l\)\" by transfer (fact gcd_code_int) lemma lcm_code_integer [code]: "lcm a b = \a\ * \b\ div gcd a b" for a b :: integer by transfer (fact lcm_altdef_int) end code_printing constant "gcd :: integer \ _" \ (OCaml) "!(fun k l -> if Z.equal k Z.zero then/ Z.abs l else if Z.equal/ l Z.zero then Z.abs k else Z.gcd k l)" and (Haskell) "Prelude.gcd" and (Scala) "_.gcd'((_)')" \ \There is no gcd operation in the SML standard library, so no code setup for SML\ text \Some code equations\ lemmas Gcd_nat_set_eq_fold [code] = Gcd_set_eq_fold [where ?'a = nat] lemmas Lcm_nat_set_eq_fold [code] = Lcm_set_eq_fold [where ?'a = nat] lemmas Gcd_int_set_eq_fold [code] = Gcd_set_eq_fold [where ?'a = int] lemmas Lcm_int_set_eq_fold [code] = Lcm_set_eq_fold [where ?'a = int] text \Fact aliases.\ lemma lcm_0_iff_nat [simp]: "lcm m n = 0 \ m = 0 \ n = 0" for m n :: nat by (fact lcm_eq_0_iff) lemma lcm_0_iff_int [simp]: "lcm m n = 0 \ m = 0 \ n = 0" for m n :: int by (fact lcm_eq_0_iff) lemma dvd_lcm_I1_nat [simp]: "k dvd m \ k dvd lcm m n" for k m n :: nat by (fact dvd_lcmI1) lemma dvd_lcm_I2_nat [simp]: "k dvd n \ k dvd lcm m n" for k m n :: nat by (fact dvd_lcmI2) lemma dvd_lcm_I1_int [simp]: "i dvd m \ i dvd lcm m n" for i m n :: int by (fact dvd_lcmI1) lemma dvd_lcm_I2_int [simp]: "i dvd n \ i dvd lcm m n" for i m n :: int by (fact dvd_lcmI2) lemmas Gcd_dvd_nat [simp] = Gcd_dvd [where ?'a = nat] lemmas Gcd_dvd_int [simp] = Gcd_dvd [where ?'a = int] lemmas Gcd_greatest_nat [simp] = Gcd_greatest [where ?'a = nat] lemmas Gcd_greatest_int [simp] = Gcd_greatest [where ?'a = int] lemma dvd_Lcm_int [simp]: "m \ M \ m dvd Lcm M" for M :: "int set" by (fact dvd_Lcm) lemma gcd_neg_numeral_1_int [simp]: "gcd (- numeral n :: int) x = gcd (numeral n) x" by (fact gcd_neg1_int) lemma gcd_neg_numeral_2_int [simp]: "gcd x (- numeral n :: int) = gcd x (numeral n)" by (fact gcd_neg2_int) lemma gcd_proj1_if_dvd_nat [simp]: "x dvd y \ gcd x y = x" for x y :: nat by (fact gcd_nat.absorb1) lemma gcd_proj2_if_dvd_nat [simp]: "y dvd x \ gcd x y = y" for x y :: nat by (fact gcd_nat.absorb2) +lemma Gcd_in: + fixes A :: "nat set" + assumes "\a b. a \ A \ b \ A \ gcd a b \ A" + assumes "A \ {}" + shows "Gcd A \ A" +proof (cases "A = {0}") + case False + with assms obtain x where "x \ A" "x > 0" + by auto + thus "Gcd A \ A" + proof (induction x rule: less_induct) + case (less x) + show ?case + proof (cases "x = Gcd A") + case False + have "\y\A. \x dvd y" + using False less.prems by (metis Gcd_dvd Gcd_greatest_nat gcd_nat.asym) + then obtain y where y: "y \ A" "\x dvd y" + by blast + have "gcd x y \ A" + by (rule assms(1)) (use \x \ A\ y in auto) + moreover have "gcd x y < x" + using \x > 0\ y by (metis gcd_dvd1 gcd_dvd2 nat_dvd_not_less nat_neq_iff) + moreover have "gcd x y > 0" + using \x > 0\ by auto + ultimately show ?thesis using less.IH by blast + qed (use less in auto) + qed +qed auto + +lemma bezout_gcd_nat': + fixes a b :: nat + shows "\x y. b * y \ a * x \ a * x - b * y = gcd a b \ a * y \ b * x \ b * x - a * y = gcd a b" + using bezout_nat[of a b] + by (metis add_diff_cancel_left' diff_zero gcd.commute gcd_0_nat + le_add_same_cancel1 mult.right_neutral zero_le) + lemmas Lcm_eq_0_I_nat [simp] = Lcm_eq_0_I [where ?'a = nat] lemmas Lcm_0_iff_nat [simp] = Lcm_0_iff [where ?'a = nat] lemmas Lcm_least_int [simp] = Lcm_least [where ?'a = int] + +subsection \Characteristic of a semiring\ + +definition (in semiring_1) semiring_char :: "'a itself \ nat" + where "semiring_char _ = Gcd {n. of_nat n = (0 :: 'a)}" + +context semiring_1 +begin + +context + fixes CHAR :: nat + defines "CHAR \ semiring_char (Pure.type :: 'a itself)" +begin + +lemma of_nat_CHAR [simp]: "of_nat CHAR = (0 :: 'a)" +proof - + have "CHAR \ {n. of_nat n = (0::'a)}" + unfolding CHAR_def semiring_char_def + proof (rule Gcd_in, clarify) + fix a b :: nat + assume *: "of_nat a = (0 :: 'a)" "of_nat b = (0 :: 'a)" + show "of_nat (gcd a b) = (0 :: 'a)" + proof (cases "a = 0") + case False + with bezout_nat obtain x y where "a * x = b * y + gcd a b" + by blast + hence "of_nat (a * x) = (of_nat (b * y + gcd a b) :: 'a)" + by (rule arg_cong) + thus "of_nat (gcd a b) = (0 :: 'a)" + using * by simp + qed (use * in auto) + next + have "of_nat 0 = (0 :: 'a)" + by simp + thus "{n. of_nat n = (0 :: 'a)} \ {}" + by blast + qed + thus ?thesis + by simp +qed + +lemma of_nat_eq_0_iff_char_dvd: "of_nat n = (0 :: 'a) \ CHAR dvd n" +proof + assume "of_nat n = (0 :: 'a)" + thus "CHAR dvd n" + unfolding CHAR_def semiring_char_def by (intro Gcd_dvd) auto +next + assume "CHAR dvd n" + then obtain m where "n = CHAR * m" + by auto + thus "of_nat n = (0 :: 'a)" + by simp +qed + +lemma CHAR_eqI: + assumes "of_nat c = (0 :: 'a)" + assumes "\x. of_nat x = (0 :: 'a) \ c dvd x" + shows "CHAR = c" + using assms by (intro dvd_antisym) (auto simp: of_nat_eq_0_iff_char_dvd) + +lemma CHAR_eq0_iff: "CHAR = 0 \ (\n>0. of_nat n \ (0::'a))" + by (auto simp: of_nat_eq_0_iff_char_dvd) + +lemma CHAR_pos_iff: "CHAR > 0 \ (\n>0. of_nat n = (0::'a))" + using CHAR_eq0_iff neq0_conv by blast + +lemma CHAR_eq_posI: + assumes "c > 0" "of_nat c = (0 :: 'a)" "\x. x > 0 \ x < c \ of_nat x \ (0 :: 'a)" + shows "CHAR = c" +proof (rule antisym) + from assms have "CHAR > 0" + by (auto simp: CHAR_pos_iff) + from assms(3)[OF this] show "CHAR \ c" + by force +next + have "CHAR dvd c" + using assms by (auto simp: of_nat_eq_0_iff_char_dvd) + thus "CHAR \ c" + using \c > 0\ by (intro dvd_imp_le) auto +qed + end +end + +lemma (in semiring_char_0) CHAR_eq_0 [simp]: "semiring_char (Pure.type :: 'a itself) = 0" + by (simp add: CHAR_eq0_iff) + + +(* nicer notation *) + +syntax "_type_char" :: "type => nat" ("(1CHAR/(1'(_')))") + +translations "CHAR('t)" => "CONST semiring_char (CONST Pure.type :: 't itself)" + +print_translation \ + let + fun char_type_tr' ctxt [Const (\<^const_syntax>\Pure.type\, Type (_, [T]))] = + Syntax.const \<^syntax_const>\_type_char\ $ Syntax_Phases.term_of_typ ctxt T + in [(\<^const_syntax>\semiring_char\, char_type_tr')] end +\ + +end diff --git a/src/HOL/Int.thy b/src/HOL/Int.thy --- a/src/HOL/Int.thy +++ b/src/HOL/Int.thy @@ -1,2246 +1,2252 @@ (* Title: HOL/Int.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Tobias Nipkow, Florian Haftmann, TU Muenchen *) section \The Integers as Equivalence Classes over Pairs of Natural Numbers\ theory Int imports Equiv_Relations Power Quotient Fun_Def begin subsection \Definition of integers as a quotient type\ definition intrel :: "(nat \ nat) \ (nat \ nat) \ bool" where "intrel = (\(x, y) (u, v). x + v = u + y)" lemma intrel_iff [simp]: "intrel (x, y) (u, v) \ x + v = u + y" by (simp add: intrel_def) quotient_type int = "nat \ nat" / "intrel" morphisms Rep_Integ Abs_Integ proof (rule equivpI) show "reflp intrel" by (auto simp: reflp_def) show "symp intrel" by (auto simp: symp_def) show "transp intrel" by (auto simp: transp_def) qed lemma eq_Abs_Integ [case_names Abs_Integ, cases type: int]: "(\x y. z = Abs_Integ (x, y) \ P) \ P" by (induct z) auto subsection \Integers form a commutative ring\ instantiation int :: comm_ring_1 begin lift_definition zero_int :: "int" is "(0, 0)" . lift_definition one_int :: "int" is "(1, 0)" . lift_definition plus_int :: "int \ int \ int" is "\(x, y) (u, v). (x + u, y + v)" by clarsimp lift_definition uminus_int :: "int \ int" is "\(x, y). (y, x)" by clarsimp lift_definition minus_int :: "int \ int \ int" is "\(x, y) (u, v). (x + v, y + u)" by clarsimp lift_definition times_int :: "int \ int \ int" is "\(x, y) (u, v). (x*u + y*v, x*v + y*u)" proof (clarsimp) fix s t u v w x y z :: nat assume "s + v = u + t" and "w + z = y + x" then have "(s + v) * w + (u + t) * x + u * (w + z) + v * (y + x) = (u + t) * w + (s + v) * x + u * (y + x) + v * (w + z)" by simp then show "(s * w + t * x) + (u * z + v * y) = (u * y + v * z) + (s * x + t * w)" by (simp add: algebra_simps) qed instance by standard (transfer; clarsimp simp: algebra_simps)+ end abbreviation int :: "nat \ int" where "int \ of_nat" lemma int_def: "int n = Abs_Integ (n, 0)" by (induct n) (simp add: zero_int.abs_eq, simp add: one_int.abs_eq plus_int.abs_eq) lemma int_transfer [transfer_rule]: includes lifting_syntax shows "rel_fun (=) pcr_int (\n. (n, 0)) int" by (simp add: rel_fun_def int.pcr_cr_eq cr_int_def int_def) lemma int_diff_cases: obtains (diff) m n where "z = int m - int n" by transfer clarsimp subsection \Integers are totally ordered\ instantiation int :: linorder begin lift_definition less_eq_int :: "int \ int \ bool" is "\(x, y) (u, v). x + v \ u + y" by auto lift_definition less_int :: "int \ int \ bool" is "\(x, y) (u, v). x + v < u + y" by auto instance by standard (transfer, force)+ end instantiation int :: distrib_lattice begin definition "(inf :: int \ int \ int) = min" definition "(sup :: int \ int \ int) = max" instance by standard (auto simp add: inf_int_def sup_int_def max_min_distrib2) end subsection \Ordering properties of arithmetic operations\ instance int :: ordered_cancel_ab_semigroup_add proof fix i j k :: int show "i \ j \ k + i \ k + j" by transfer clarsimp qed text \Strict Monotonicity of Multiplication.\ text \Strict, in 1st argument; proof is by induction on \k > 0\.\ lemma zmult_zless_mono2_lemma: "i < j \ 0 < k \ int k * i < int k * j" for i j :: int proof (induct k) case 0 then show ?case by simp next case (Suc k) then show ?case by (cases "k = 0") (simp_all add: distrib_right add_strict_mono) qed lemma zero_le_imp_eq_int: assumes "k \ (0::int)" shows "\n. k = int n" proof - have "b \ a \ \n::nat. a = n + b" for a b by (rule_tac x="a - b" in exI) simp with assms show ?thesis by transfer auto qed lemma zero_less_imp_eq_int: assumes "k > (0::int)" shows "\n>0. k = int n" proof - have "b < a \ \n::nat. n>0 \ a = n + b" for a b by (rule_tac x="a - b" in exI) simp with assms show ?thesis by transfer auto qed lemma zmult_zless_mono2: "i < j \ 0 < k \ k * i < k * j" for i j k :: int by (drule zero_less_imp_eq_int) (auto simp add: zmult_zless_mono2_lemma) text \The integers form an ordered integral domain.\ instantiation int :: linordered_idom begin definition zabs_def: "\i::int\ = (if i < 0 then - i else i)" definition zsgn_def: "sgn (i::int) = (if i = 0 then 0 else if 0 < i then 1 else - 1)" instance proof fix i j k :: int show "i < j \ 0 < k \ k * i < k * j" by (rule zmult_zless_mono2) show "\i\ = (if i < 0 then -i else i)" by (simp only: zabs_def) show "sgn (i::int) = (if i=0 then 0 else if 0 w + 1 \ z" for w z :: int by transfer clarsimp lemma zless_iff_Suc_zadd: "w < z \ (\n. z = w + int (Suc n))" for w z :: int proof - have "\a b c d. a + d < c + b \ \n. c + b = Suc (a + n + d)" by (rule_tac x="c+b - Suc(a+d)" in exI) arith then show ?thesis by transfer auto qed lemma zabs_less_one_iff [simp]: "\z\ < 1 \ z = 0" (is "?lhs \ ?rhs") for z :: int proof assume ?rhs then show ?lhs by simp next assume ?lhs with zless_imp_add1_zle [of "\z\" 1] have "\z\ + 1 \ 1" by simp then have "\z\ \ 0" by simp then show ?rhs by simp qed subsection \Embedding of the Integers into any \ring_1\: \of_int\\ context ring_1 begin lift_definition of_int :: "int \ 'a" is "\(i, j). of_nat i - of_nat j" by (clarsimp simp add: diff_eq_eq eq_diff_eq diff_add_eq of_nat_add [symmetric] simp del: of_nat_add) lemma of_int_0 [simp]: "of_int 0 = 0" by transfer simp lemma of_int_1 [simp]: "of_int 1 = 1" by transfer simp lemma of_int_add [simp]: "of_int (w + z) = of_int w + of_int z" by transfer (clarsimp simp add: algebra_simps) lemma of_int_minus [simp]: "of_int (- z) = - (of_int z)" by (transfer fixing: uminus) clarsimp lemma of_int_diff [simp]: "of_int (w - z) = of_int w - of_int z" using of_int_add [of w "- z"] by simp lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z" by (transfer fixing: times) (clarsimp simp add: algebra_simps) lemma mult_of_int_commute: "of_int x * y = y * of_int x" by (transfer fixing: times) (auto simp: algebra_simps mult_of_nat_commute) text \Collapse nested embeddings.\ lemma of_int_of_nat_eq [simp]: "of_int (int n) = of_nat n" by (induct n) auto lemma of_int_numeral [simp, code_post]: "of_int (numeral k) = numeral k" by (simp add: of_nat_numeral [symmetric] of_int_of_nat_eq [symmetric]) lemma of_int_neg_numeral [code_post]: "of_int (- numeral k) = - numeral k" by simp lemma of_int_power [simp]: "of_int (z ^ n) = of_int z ^ n" by (induct n) simp_all lemma of_int_of_bool [simp]: "of_int (of_bool P) = of_bool P" by auto end context ring_char_0 begin lemma of_int_eq_iff [simp]: "of_int w = of_int z \ w = z" by transfer (clarsimp simp add: algebra_simps of_nat_add [symmetric] simp del: of_nat_add) text \Special cases where either operand is zero.\ lemma of_int_eq_0_iff [simp]: "of_int z = 0 \ z = 0" using of_int_eq_iff [of z 0] by simp lemma of_int_0_eq_iff [simp]: "0 = of_int z \ z = 0" using of_int_eq_iff [of 0 z] by simp lemma of_int_eq_1_iff [iff]: "of_int z = 1 \ z = 1" using of_int_eq_iff [of z 1] by simp lemma numeral_power_eq_of_int_cancel_iff [simp]: "numeral x ^ n = of_int y \ numeral x ^ n = y" using of_int_eq_iff[of "numeral x ^ n" y, unfolded of_int_numeral of_int_power] . lemma of_int_eq_numeral_power_cancel_iff [simp]: "of_int y = numeral x ^ n \ y = numeral x ^ n" using numeral_power_eq_of_int_cancel_iff [of x n y] by (metis (mono_tags)) lemma neg_numeral_power_eq_of_int_cancel_iff [simp]: "(- numeral x) ^ n = of_int y \ (- numeral x) ^ n = y" using of_int_eq_iff[of "(- numeral x) ^ n" y] by simp lemma of_int_eq_neg_numeral_power_cancel_iff [simp]: "of_int y = (- numeral x) ^ n \ y = (- numeral x) ^ n" using neg_numeral_power_eq_of_int_cancel_iff[of x n y] by (metis (mono_tags)) lemma of_int_eq_of_int_power_cancel_iff[simp]: "(of_int b) ^ w = of_int x \ b ^ w = x" by (metis of_int_power of_int_eq_iff) lemma of_int_power_eq_of_int_cancel_iff[simp]: "of_int x = (of_int b) ^ w \ x = b ^ w" by (metis of_int_eq_of_int_power_cancel_iff) end context linordered_idom begin text \Every \linordered_idom\ has characteristic zero.\ subclass ring_char_0 .. lemma of_int_le_iff [simp]: "of_int w \ of_int z \ w \ z" by (transfer fixing: less_eq) (clarsimp simp add: algebra_simps of_nat_add [symmetric] simp del: of_nat_add) lemma of_int_less_iff [simp]: "of_int w < of_int z \ w < z" by (simp add: less_le order_less_le) lemma of_int_0_le_iff [simp]: "0 \ of_int z \ 0 \ z" using of_int_le_iff [of 0 z] by simp lemma of_int_le_0_iff [simp]: "of_int z \ 0 \ z \ 0" using of_int_le_iff [of z 0] by simp lemma of_int_0_less_iff [simp]: "0 < of_int z \ 0 < z" using of_int_less_iff [of 0 z] by simp lemma of_int_less_0_iff [simp]: "of_int z < 0 \ z < 0" using of_int_less_iff [of z 0] by simp lemma of_int_1_le_iff [simp]: "1 \ of_int z \ 1 \ z" using of_int_le_iff [of 1 z] by simp lemma of_int_le_1_iff [simp]: "of_int z \ 1 \ z \ 1" using of_int_le_iff [of z 1] by simp lemma of_int_1_less_iff [simp]: "1 < of_int z \ 1 < z" using of_int_less_iff [of 1 z] by simp lemma of_int_less_1_iff [simp]: "of_int z < 1 \ z < 1" using of_int_less_iff [of z 1] by simp lemma of_int_pos: "z > 0 \ of_int z > 0" by simp lemma of_int_nonneg: "z \ 0 \ of_int z \ 0" by simp lemma of_int_abs [simp]: "of_int \x\ = \of_int x\" by (auto simp add: abs_if) lemma of_int_lessD: assumes "\of_int n\ < x" shows "n = 0 \ x > 1" proof (cases "n = 0") case True then show ?thesis by simp next case False then have "\n\ \ 0" by simp then have "\n\ > 0" by simp then have "\n\ \ 1" using zless_imp_add1_zle [of 0 "\n\"] by simp then have "\of_int n\ \ 1" unfolding of_int_1_le_iff [of "\n\", symmetric] by simp then have "1 < x" using assms by (rule le_less_trans) then show ?thesis .. qed lemma of_int_leD: assumes "\of_int n\ \ x" shows "n = 0 \ 1 \ x" proof (cases "n = 0") case True then show ?thesis by simp next case False then have "\n\ \ 0" by simp then have "\n\ > 0" by simp then have "\n\ \ 1" using zless_imp_add1_zle [of 0 "\n\"] by simp then have "\of_int n\ \ 1" unfolding of_int_1_le_iff [of "\n\", symmetric] by simp then have "1 \ x" using assms by (rule order_trans) then show ?thesis .. qed lemma numeral_power_le_of_int_cancel_iff [simp]: "numeral x ^ n \ of_int a \ numeral x ^ n \ a" by (metis (mono_tags) local.of_int_eq_numeral_power_cancel_iff of_int_le_iff) lemma of_int_le_numeral_power_cancel_iff [simp]: "of_int a \ numeral x ^ n \ a \ numeral x ^ n" by (metis (mono_tags) local.numeral_power_eq_of_int_cancel_iff of_int_le_iff) lemma numeral_power_less_of_int_cancel_iff [simp]: "numeral x ^ n < of_int a \ numeral x ^ n < a" by (metis (mono_tags) local.of_int_eq_numeral_power_cancel_iff of_int_less_iff) lemma of_int_less_numeral_power_cancel_iff [simp]: "of_int a < numeral x ^ n \ a < numeral x ^ n" by (metis (mono_tags) local.of_int_eq_numeral_power_cancel_iff of_int_less_iff) lemma neg_numeral_power_le_of_int_cancel_iff [simp]: "(- numeral x) ^ n \ of_int a \ (- numeral x) ^ n \ a" by (metis (mono_tags) of_int_le_iff of_int_neg_numeral of_int_power) lemma of_int_le_neg_numeral_power_cancel_iff [simp]: "of_int a \ (- numeral x) ^ n \ a \ (- numeral x) ^ n" by (metis (mono_tags) of_int_le_iff of_int_neg_numeral of_int_power) lemma neg_numeral_power_less_of_int_cancel_iff [simp]: "(- numeral x) ^ n < of_int a \ (- numeral x) ^ n < a" using of_int_less_iff[of "(- numeral x) ^ n" a] by simp lemma of_int_less_neg_numeral_power_cancel_iff [simp]: "of_int a < (- numeral x) ^ n \ a < (- numeral x::int) ^ n" using of_int_less_iff[of a "(- numeral x) ^ n"] by simp lemma of_int_le_of_int_power_cancel_iff[simp]: "(of_int b) ^ w \ of_int x \ b ^ w \ x" by (metis (mono_tags) of_int_le_iff of_int_power) lemma of_int_power_le_of_int_cancel_iff[simp]: "of_int x \ (of_int b) ^ w\ x \ b ^ w" by (metis (mono_tags) of_int_le_iff of_int_power) lemma of_int_less_of_int_power_cancel_iff[simp]: "(of_int b) ^ w < of_int x \ b ^ w < x" by (metis (mono_tags) of_int_less_iff of_int_power) lemma of_int_power_less_of_int_cancel_iff[simp]: "of_int x < (of_int b) ^ w\ x < b ^ w" by (metis (mono_tags) of_int_less_iff of_int_power) lemma of_int_max: "of_int (max x y) = max (of_int x) (of_int y)" by (auto simp: max_def) lemma of_int_min: "of_int (min x y) = min (of_int x) (of_int y)" by (auto simp: min_def) end context division_ring begin lemmas mult_inverse_of_int_commute = mult_commute_imp_mult_inverse_commute[OF mult_of_int_commute] end text \Comparisons involving \<^term>\of_int\.\ lemma of_int_eq_numeral_iff [iff]: "of_int z = (numeral n :: 'a::ring_char_0) \ z = numeral n" using of_int_eq_iff by fastforce lemma of_int_le_numeral_iff [simp]: "of_int z \ (numeral n :: 'a::linordered_idom) \ z \ numeral n" using of_int_le_iff [of z "numeral n"] by simp lemma of_int_numeral_le_iff [simp]: "(numeral n :: 'a::linordered_idom) \ of_int z \ numeral n \ z" using of_int_le_iff [of "numeral n"] by simp lemma of_int_less_numeral_iff [simp]: "of_int z < (numeral n :: 'a::linordered_idom) \ z < numeral n" using of_int_less_iff [of z "numeral n"] by simp lemma of_int_numeral_less_iff [simp]: "(numeral n :: 'a::linordered_idom) < of_int z \ numeral n < z" using of_int_less_iff [of "numeral n" z] by simp lemma of_nat_less_of_int_iff: "(of_nat n::'a::linordered_idom) < of_int x \ int n < x" by (metis of_int_of_nat_eq of_int_less_iff) lemma of_int_eq_id [simp]: "of_int = id" proof show "of_int z = id z" for z by (cases z rule: int_diff_cases) simp qed instance int :: no_top proof show "\x::int. \y. x < y" by (rule_tac x="x + 1" in exI) simp qed instance int :: no_bot proof show "\x::int. \y. y < x" by (rule_tac x="x - 1" in exI) simp qed subsection \Magnitude of an Integer, as a Natural Number: \nat\\ lift_definition nat :: "int \ nat" is "\(x, y). x - y" by auto lemma nat_int [simp]: "nat (int n) = n" by transfer simp lemma int_nat_eq [simp]: "int (nat z) = (if 0 \ z then z else 0)" by transfer clarsimp lemma nat_0_le: "0 \ z \ int (nat z) = z" by simp lemma nat_le_0 [simp]: "z \ 0 \ nat z = 0" by transfer clarsimp lemma nat_le_eq_zle: "0 < w \ 0 \ z \ nat w \ nat z \ w \ z" by transfer (clarsimp, arith) text \An alternative condition is \<^term>\0 \ w\.\ lemma nat_mono_iff: "0 < z \ nat w < nat z \ w < z" by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) lemma nat_less_eq_zless: "0 \ w \ nat w < nat z \ w < z" by (simp add: nat_le_eq_zle linorder_not_le [symmetric]) lemma zless_nat_conj [simp]: "nat w < nat z \ 0 < z \ w < z" by transfer (clarsimp, arith) lemma nonneg_int_cases: assumes "0 \ k" obtains n where "k = int n" proof - from assms have "k = int (nat k)" by simp then show thesis by (rule that) qed lemma pos_int_cases: assumes "0 < k" obtains n where "k = int n" and "n > 0" proof - from assms have "0 \ k" by simp then obtain n where "k = int n" by (rule nonneg_int_cases) moreover have "n > 0" using \k = int n\ assms by simp ultimately show thesis by (rule that) qed lemma nonpos_int_cases: assumes "k \ 0" obtains n where "k = - int n" proof - from assms have "- k \ 0" by simp then obtain n where "- k = int n" by (rule nonneg_int_cases) then have "k = - int n" by simp then show thesis by (rule that) qed lemma neg_int_cases: assumes "k < 0" obtains n where "k = - int n" and "n > 0" proof - from assms have "- k > 0" by simp then obtain n where "- k = int n" and "- k > 0" by (blast elim: pos_int_cases) then have "k = - int n" and "n > 0" by simp_all then show thesis by (rule that) qed lemma nat_eq_iff: "nat w = m \ (if 0 \ w then w = int m else m = 0)" by transfer (clarsimp simp add: le_imp_diff_is_add) lemma nat_eq_iff2: "m = nat w \ (if 0 \ w then w = int m else m = 0)" using nat_eq_iff [of w m] by auto lemma nat_0 [simp]: "nat 0 = 0" by (simp add: nat_eq_iff) lemma nat_1 [simp]: "nat 1 = Suc 0" by (simp add: nat_eq_iff) lemma nat_numeral [simp]: "nat (numeral k) = numeral k" by (simp add: nat_eq_iff) lemma nat_neg_numeral [simp]: "nat (- numeral k) = 0" by simp lemma nat_2: "nat 2 = Suc (Suc 0)" by simp lemma nat_less_iff: "0 \ w \ nat w < m \ w < of_nat m" by transfer (clarsimp, arith) lemma nat_le_iff: "nat x \ n \ x \ int n" by transfer (clarsimp simp add: le_diff_conv) lemma nat_mono: "x \ y \ nat x \ nat y" by transfer auto lemma nat_0_iff[simp]: "nat i = 0 \ i \ 0" for i :: int by transfer clarsimp lemma int_eq_iff: "of_nat m = z \ m = nat z \ 0 \ z" by (auto simp add: nat_eq_iff2) lemma zero_less_nat_eq [simp]: "0 < nat z \ 0 < z" using zless_nat_conj [of 0] by auto lemma nat_add_distrib: "0 \ z \ 0 \ z' \ nat (z + z') = nat z + nat z'" by transfer clarsimp lemma nat_diff_distrib': "0 \ x \ 0 \ y \ nat (x - y) = nat x - nat y" by transfer clarsimp lemma nat_diff_distrib: "0 \ z' \ z' \ z \ nat (z - z') = nat z - nat z'" by (rule nat_diff_distrib') auto lemma nat_zminus_int [simp]: "nat (- int n) = 0" by transfer simp lemma le_nat_iff: "k \ 0 \ n \ nat k \ int n \ k" by transfer auto lemma zless_nat_eq_int_zless: "m < nat z \ int m < z" by transfer (clarsimp simp add: less_diff_conv) lemma (in ring_1) of_nat_nat [simp]: "0 \ z \ of_nat (nat z) = of_int z" by transfer (clarsimp simp add: of_nat_diff) lemma diff_nat_numeral [simp]: "(numeral v :: nat) - numeral v' = nat (numeral v - numeral v')" by (simp only: nat_diff_distrib' zero_le_numeral nat_numeral) lemma nat_abs_triangle_ineq: "nat \k + l\ \ nat \k\ + nat \l\" by (simp add: nat_add_distrib [symmetric] nat_le_eq_zle abs_triangle_ineq) lemma nat_of_bool [simp]: "nat (of_bool P) = of_bool P" by auto lemma split_nat [arith_split]: "P (nat i) \ ((\n. i = int n \ P n) \ (i < 0 \ P 0))" (is "?P = (?L \ ?R)") for i :: int proof (cases "i < 0") case True then show ?thesis by auto next case False have "?P = ?L" proof assume ?P then show ?L using False by auto next assume ?L moreover from False have "int (nat i) = i" by (simp add: not_less) ultimately show ?P by simp qed with False show ?thesis by simp qed lemma all_nat: "(\x. P x) \ (\x\0. P (nat x))" by (auto split: split_nat) lemma ex_nat: "(\x. P x) \ (\x. 0 \ x \ P (nat x))" proof assume "\x. P x" then obtain x where "P x" .. then have "int x \ 0 \ P (nat (int x))" by simp then show "\x\0. P (nat x)" .. next assume "\x\0. P (nat x)" then show "\x. P x" by auto qed text \For termination proofs:\ lemma measure_function_int[measure_function]: "is_measure (nat \ abs)" .. subsection \Lemmas about the Function \<^term>\of_nat\ and Orderings\ lemma negative_zless_0: "- (int (Suc n)) < (0 :: int)" by (simp add: order_less_le del: of_nat_Suc) lemma negative_zless [iff]: "- (int (Suc n)) < int m" by (rule negative_zless_0 [THEN order_less_le_trans], simp) lemma negative_zle_0: "- int n \ 0" by (simp add: minus_le_iff) lemma negative_zle [iff]: "- int n \ int m" by (rule order_trans [OF negative_zle_0 of_nat_0_le_iff]) lemma not_zle_0_negative [simp]: "\ 0 \ - int (Suc n)" by (subst le_minus_iff) (simp del: of_nat_Suc) lemma int_zle_neg: "int n \ - int m \ n = 0 \ m = 0" by transfer simp lemma not_int_zless_negative [simp]: "\ int n < - int m" by (simp add: linorder_not_less) lemma negative_eq_positive [simp]: "- int n = of_nat m \ n = 0 \ m = 0" by (force simp add: order_eq_iff [of "- of_nat n"] int_zle_neg) lemma zle_iff_zadd: "w \ z \ (\n. z = w + int n)" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by auto next assume ?lhs then have "0 \ z - w" by simp then obtain n where "z - w = int n" using zero_le_imp_eq_int [of "z - w"] by blast then have "z = w + int n" by simp then show ?rhs .. qed lemma zadd_int_left: "int m + (int n + z) = int (m + n) + z" by simp text \ This version is proved for all ordered rings, not just integers! It is proved here because attribute \arith_split\ is not available in theory \Rings\. But is it really better than just rewriting with \abs_if\? \ lemma abs_split [arith_split, no_atp]: "P \a\ \ (0 \ a \ P a) \ (a < 0 \ P (- a))" for a :: "'a::linordered_idom" by (force dest: order_less_le_trans simp add: abs_if linorder_not_less) lemma negD: assumes "x < 0" shows "\n. x = - (int (Suc n))" proof - have "\a b. a < b \ \n. Suc (a + n) = b" by (rule_tac x="b - Suc a" in exI) arith with assms show ?thesis by transfer auto qed subsection \Cases and induction\ text \ Now we replace the case analysis rule by a more conventional one: whether an integer is negative or not. \ text \This version is symmetric in the two subgoals.\ lemma int_cases2 [case_names nonneg nonpos, cases type: int]: "(\n. z = int n \ P) \ (\n. z = - (int n) \ P) \ P" by (cases "z < 0") (auto simp add: linorder_not_less dest!: negD nat_0_le [THEN sym]) text \This is the default, with a negative case.\ lemma int_cases [case_names nonneg neg, cases type: int]: assumes pos: "\n. z = int n \ P" and neg: "\n. z = - (int (Suc n)) \ P" shows P proof (cases "z < 0") case True with neg show ?thesis by (blast dest!: negD) next case False with pos show ?thesis by (force simp add: linorder_not_less dest: nat_0_le [THEN sym]) qed lemma int_cases3 [case_names zero pos neg]: fixes k :: int assumes "k = 0 \ P" and "\n. k = int n \ n > 0 \ P" and "\n. k = - int n \ n > 0 \ P" shows "P" proof (cases k "0::int" rule: linorder_cases) case equal with assms(1) show P by simp next case greater then have *: "nat k > 0" by simp moreover from * have "k = int (nat k)" by auto ultimately show P using assms(2) by blast next case less then have *: "nat (- k) > 0" by simp moreover from * have "k = - int (nat (- k))" by auto ultimately show P using assms(3) by blast qed lemma int_of_nat_induct [case_names nonneg neg, induct type: int]: "(\n. P (int n)) \ (\n. P (- (int (Suc n)))) \ P z" by (cases z) auto lemma sgn_mult_dvd_iff [simp]: "sgn r * l dvd k \ l dvd k \ (r = 0 \ k = 0)" for k l r :: int by (cases r rule: int_cases3) auto lemma mult_sgn_dvd_iff [simp]: "l * sgn r dvd k \ l dvd k \ (r = 0 \ k = 0)" for k l r :: int using sgn_mult_dvd_iff [of r l k] by (simp add: ac_simps) lemma dvd_sgn_mult_iff [simp]: "l dvd sgn r * k \ l dvd k \ r = 0" for k l r :: int by (cases r rule: int_cases3) simp_all lemma dvd_mult_sgn_iff [simp]: "l dvd k * sgn r \ l dvd k \ r = 0" for k l r :: int using dvd_sgn_mult_iff [of l r k] by (simp add: ac_simps) lemma int_sgnE: fixes k :: int obtains n and l where "k = sgn l * int n" proof - have "k = sgn k * int (nat \k\)" by (simp add: sgn_mult_abs) then show ?thesis .. qed subsubsection \Binary comparisons\ text \Preliminaries\ lemma le_imp_0_less: fixes z :: int assumes le: "0 \ z" shows "0 < 1 + z" proof - have "0 \ z" by fact also have "\ < z + 1" by (rule less_add_one) also have "\ = 1 + z" by (simp add: ac_simps) finally show "0 < 1 + z" . qed lemma odd_less_0_iff: "1 + z + z < 0 \ z < 0" for z :: int proof (cases z) case (nonneg n) then show ?thesis by (simp add: linorder_not_less add.assoc add_increasing le_imp_0_less [THEN order_less_imp_le]) next case (neg n) then show ?thesis by (simp del: of_nat_Suc of_nat_add of_nat_1 add: algebra_simps of_nat_1 [where 'a=int, symmetric] of_nat_add [symmetric]) qed subsubsection \Comparisons, for Ordered Rings\ lemma odd_nonzero: "1 + z + z \ 0" for z :: int proof (cases z) case (nonneg n) have le: "0 \ z + z" by (simp add: nonneg add_increasing) then show ?thesis using le_imp_0_less [OF le] by (auto simp: ac_simps) next case (neg n) show ?thesis proof assume eq: "1 + z + z = 0" have "0 < 1 + (int n + int n)" by (simp add: le_imp_0_less add_increasing) also have "\ = - (1 + z + z)" by (simp add: neg add.assoc [symmetric]) also have "\ = 0" by (simp add: eq) finally have "0<0" .. then show False by blast qed qed subsection \The Set of Integers\ context ring_1 begin definition Ints :: "'a set" ("\") where "\ = range of_int" lemma Ints_of_int [simp]: "of_int z \ \" by (simp add: Ints_def) lemma Ints_of_nat [simp]: "of_nat n \ \" using Ints_of_int [of "of_nat n"] by simp lemma Ints_0 [simp]: "0 \ \" using Ints_of_int [of "0"] by simp lemma Ints_1 [simp]: "1 \ \" using Ints_of_int [of "1"] by simp lemma Ints_numeral [simp]: "numeral n \ \" by (subst of_nat_numeral [symmetric], rule Ints_of_nat) lemma Ints_add [simp]: "a \ \ \ b \ \ \ a + b \ \" by (force simp add: Ints_def simp flip: of_int_add intro: range_eqI) lemma Ints_minus [simp]: "a \ \ \ -a \ \" by (force simp add: Ints_def simp flip: of_int_minus intro: range_eqI) lemma minus_in_Ints_iff: "-x \ \ \ x \ \" using Ints_minus[of x] Ints_minus[of "-x"] by auto lemma Ints_diff [simp]: "a \ \ \ b \ \ \ a - b \ \" by (force simp add: Ints_def simp flip: of_int_diff intro: range_eqI) lemma Ints_mult [simp]: "a \ \ \ b \ \ \ a * b \ \" by (force simp add: Ints_def simp flip: of_int_mult intro: range_eqI) lemma Ints_power [simp]: "a \ \ \ a ^ n \ \" by (induct n) simp_all lemma Ints_cases [cases set: Ints]: assumes "q \ \" obtains (of_int) z where "q = of_int z" unfolding Ints_def proof - from \q \ \\ have "q \ range of_int" unfolding Ints_def . then obtain z where "q = of_int z" .. then show thesis .. qed lemma Ints_induct [case_names of_int, induct set: Ints]: "q \ \ \ (\z. P (of_int z)) \ P q" by (rule Ints_cases) auto lemma Nats_subset_Ints: "\ \ \" unfolding Nats_def Ints_def by (rule subsetI, elim imageE, hypsubst, subst of_int_of_nat_eq[symmetric], rule imageI) simp_all lemma Nats_altdef1: "\ = {of_int n |n. n \ 0}" proof (intro subsetI equalityI) fix x :: 'a assume "x \ {of_int n |n. n \ 0}" then obtain n where "x = of_int n" "n \ 0" by (auto elim!: Ints_cases) then have "x = of_nat (nat n)" by (subst of_nat_nat) simp_all then show "x \ \" by simp next fix x :: 'a assume "x \ \" then obtain n where "x = of_nat n" by (auto elim!: Nats_cases) then have "x = of_int (int n)" by simp also have "int n \ 0" by simp then have "of_int (int n) \ {of_int n |n. n \ 0}" by blast finally show "x \ {of_int n |n. n \ 0}" . qed end +lemma Ints_sum [intro]: "(\x. x \ A \ f x \ \) \ sum f A \ \" + by (induction A rule: infinite_finite_induct) auto + +lemma Ints_prod [intro]: "(\x. x \ A \ f x \ \) \ prod f A \ \" + by (induction A rule: infinite_finite_induct) auto + lemma (in linordered_idom) Ints_abs [simp]: shows "a \ \ \ abs a \ \" by (auto simp: abs_if) lemma (in linordered_idom) Nats_altdef2: "\ = {n \ \. n \ 0}" proof (intro subsetI equalityI) fix x :: 'a assume "x \ {n \ \. n \ 0}" then obtain n where "x = of_int n" "n \ 0" by (auto elim!: Ints_cases) then have "x = of_nat (nat n)" by (subst of_nat_nat) simp_all then show "x \ \" by simp qed (auto elim!: Nats_cases) lemma (in idom_divide) of_int_divide_in_Ints: "of_int a div of_int b \ \" if "b dvd a" proof - from that obtain c where "a = b * c" .. then show ?thesis by (cases "of_int b = 0") simp_all qed text \The premise involving \<^term>\Ints\ prevents \<^term>\a = 1/2\.\ lemma Ints_double_eq_0_iff: fixes a :: "'a::ring_char_0" assumes in_Ints: "a \ \" shows "a + a = 0 \ a = 0" (is "?lhs \ ?rhs") proof - from in_Ints have "a \ range of_int" unfolding Ints_def [symmetric] . then obtain z where a: "a = of_int z" .. show ?thesis proof assume ?rhs then show ?lhs by simp next assume ?lhs with a have "of_int (z + z) = (of_int 0 :: 'a)" by simp then have "z + z = 0" by (simp only: of_int_eq_iff) then have "z = 0" by (simp only: double_zero) with a show ?rhs by simp qed qed lemma Ints_odd_nonzero: fixes a :: "'a::ring_char_0" assumes in_Ints: "a \ \" shows "1 + a + a \ 0" proof - from in_Ints have "a \ range of_int" unfolding Ints_def [symmetric] . then obtain z where a: "a = of_int z" .. show ?thesis proof assume "1 + a + a = 0" with a have "of_int (1 + z + z) = (of_int 0 :: 'a)" by simp then have "1 + z + z = 0" by (simp only: of_int_eq_iff) with odd_nonzero show False by blast qed qed lemma Nats_numeral [simp]: "numeral w \ \" using of_nat_in_Nats [of "numeral w"] by simp lemma Ints_odd_less_0: fixes a :: "'a::linordered_idom" assumes in_Ints: "a \ \" shows "1 + a + a < 0 \ a < 0" proof - from in_Ints have "a \ range of_int" unfolding Ints_def [symmetric] . then obtain z where a: "a = of_int z" .. with a have "1 + a + a < 0 \ of_int (1 + z + z) < (of_int 0 :: 'a)" by simp also have "\ \ z < 0" by (simp only: of_int_less_iff odd_less_0_iff) also have "\ \ a < 0" by (simp add: a) finally show ?thesis . qed subsection \\<^term>\sum\ and \<^term>\prod\\ context semiring_1 begin lemma of_nat_sum [simp]: "of_nat (sum f A) = (\x\A. of_nat (f x))" by (induction A rule: infinite_finite_induct) auto end context ring_1 begin lemma of_int_sum [simp]: "of_int (sum f A) = (\x\A. of_int (f x))" by (induction A rule: infinite_finite_induct) auto end context comm_semiring_1 begin lemma of_nat_prod [simp]: "of_nat (prod f A) = (\x\A. of_nat (f x))" by (induction A rule: infinite_finite_induct) auto end context comm_ring_1 begin lemma of_int_prod [simp]: "of_int (prod f A) = (\x\A. of_int (f x))" by (induction A rule: infinite_finite_induct) auto end subsection \Setting up simplification procedures\ ML_file \Tools/int_arith.ML\ declaration \K ( Lin_Arith.add_discrete_type \<^type_name>\Int.int\ #> Lin_Arith.add_lessD @{thm zless_imp_add1_zle} #> Lin_Arith.add_inj_thms @{thms of_nat_le_iff [THEN iffD2] of_nat_eq_iff [THEN iffD2]} #> Lin_Arith.add_inj_const (\<^const_name>\of_nat\, \<^typ>\nat \ int\) #> Lin_Arith.add_simps @{thms of_int_0 of_int_1 of_int_add of_int_mult of_int_numeral of_int_neg_numeral nat_0 nat_1 diff_nat_numeral nat_numeral neg_less_iff_less True_implies_equals distrib_left [where a = "numeral v" for v] distrib_left [where a = "- numeral v" for v] div_by_1 div_0 times_divide_eq_right times_divide_eq_left minus_divide_left [THEN sym] minus_divide_right [THEN sym] add_divide_distrib diff_divide_distrib of_int_minus of_int_diff of_int_of_nat_eq} #> Lin_Arith.add_simprocs [Int_Arith.zero_one_idom_simproc] )\ simproc_setup fast_arith ("(m::'a::linordered_idom) < n" | "(m::'a::linordered_idom) \ n" | "(m::'a::linordered_idom) = n") = \K Lin_Arith.simproc\ subsection\More Inequality Reasoning\ lemma zless_add1_eq: "w < z + 1 \ w < z \ w = z" for w z :: int by arith lemma add1_zle_eq: "w + 1 \ z \ w < z" for w z :: int by arith lemma zle_diff1_eq [simp]: "w \ z - 1 \ w < z" for w z :: int by arith lemma zle_add1_eq_le [simp]: "w < z + 1 \ w \ z" for w z :: int by arith lemma int_one_le_iff_zero_less: "1 \ z \ 0 < z" for z :: int by arith lemma Ints_nonzero_abs_ge1: fixes x:: "'a :: linordered_idom" assumes "x \ Ints" "x \ 0" shows "1 \ abs x" proof (rule Ints_cases [OF \x \ Ints\]) fix z::int assume "x = of_int z" with \x \ 0\ show "1 \ \x\" apply (auto simp add: abs_if) by (metis diff_0 of_int_1 of_int_le_iff of_int_minus zle_diff1_eq) qed lemma Ints_nonzero_abs_less1: fixes x:: "'a :: linordered_idom" shows "\x \ Ints; abs x < 1\ \ x = 0" using Ints_nonzero_abs_ge1 [of x] by auto lemma Ints_eq_abs_less1: fixes x:: "'a :: linordered_idom" shows "\x \ Ints; y \ Ints\ \ x = y \ abs (x-y) < 1" using eq_iff_diff_eq_0 by (fastforce intro: Ints_nonzero_abs_less1) subsection \The functions \<^term>\nat\ and \<^term>\int\\ text \Simplify the term \<^term>\w + - z\.\ lemma one_less_nat_eq [simp]: "Suc 0 < nat z \ 1 < z" using zless_nat_conj [of 1 z] by auto lemma int_eq_iff_numeral [simp]: "int m = numeral v \ m = numeral v" by (simp add: int_eq_iff) lemma nat_abs_int_diff: "nat \int a - int b\ = (if a \ b then b - a else a - b)" by auto lemma nat_int_add: "nat (int a + int b) = a + b" by auto context ring_1 begin lemma of_int_of_nat [nitpick_simp]: "of_int k = (if k < 0 then - of_nat (nat (- k)) else of_nat (nat k))" proof (cases "k < 0") case True then have "0 \ - k" by simp then have "of_nat (nat (- k)) = of_int (- k)" by (rule of_nat_nat) with True show ?thesis by simp next case False then show ?thesis by (simp add: not_less) qed end lemma transfer_rule_of_int: includes lifting_syntax fixes R :: "'a::ring_1 \ 'b::ring_1 \ bool" assumes [transfer_rule]: "R 0 0" "R 1 1" "(R ===> R ===> R) (+) (+)" "(R ===> R) uminus uminus" shows "((=) ===> R) of_int of_int" proof - note assms note transfer_rule_of_nat [transfer_rule] have [transfer_rule]: "((=) ===> R) of_nat of_nat" by transfer_prover show ?thesis by (unfold of_int_of_nat [abs_def]) transfer_prover qed lemma nat_mult_distrib: fixes z z' :: int assumes "0 \ z" shows "nat (z * z') = nat z * nat z'" proof (cases "0 \ z'") case False with assms have "z * z' \ 0" by (simp add: not_le mult_le_0_iff) then have "nat (z * z') = 0" by simp moreover from False have "nat z' = 0" by simp ultimately show ?thesis by simp next case True with assms have ge_0: "z * z' \ 0" by (simp add: zero_le_mult_iff) show ?thesis by (rule injD [of "of_nat :: nat \ int", OF inj_of_nat]) (simp only: of_nat_mult of_nat_nat [OF True] of_nat_nat [OF assms] of_nat_nat [OF ge_0], simp) qed lemma nat_mult_distrib_neg: assumes "z \ (0::int)" shows "nat (z * z') = nat (- z) * nat (- z')" (is "?L = ?R") proof - have "?L = nat (- z * - z')" using assms by auto also have "... = ?R" by (rule nat_mult_distrib) (use assms in auto) finally show ?thesis . qed lemma nat_abs_mult_distrib: "nat \w * z\ = nat \w\ * nat \z\" by (cases "z = 0 \ w = 0") (auto simp add: abs_if nat_mult_distrib [symmetric] nat_mult_distrib_neg [symmetric] mult_less_0_iff) lemma int_in_range_abs [simp]: "int n \ range abs" proof (rule range_eqI) show "int n = \int n\" by simp qed lemma range_abs_Nats [simp]: "range abs = (\ :: int set)" proof - have "\k\ \ \" for k :: int by (cases k) simp_all moreover have "k \ range abs" if "k \ \" for k :: int using that by induct simp ultimately show ?thesis by blast qed lemma Suc_nat_eq_nat_zadd1: "0 \ z \ Suc (nat z) = nat (1 + z)" for z :: int by (rule sym) (simp add: nat_eq_iff) lemma diff_nat_eq_if: "nat z - nat z' = (if z' < 0 then nat z else let d = z - z' in if d < 0 then 0 else nat d)" by (simp add: Let_def nat_diff_distrib [symmetric]) lemma nat_numeral_diff_1 [simp]: "numeral v - (1::nat) = nat (numeral v - 1)" using diff_nat_numeral [of v Num.One] by simp subsection \Induction principles for int\ text \Well-founded segments of the integers.\ definition int_ge_less_than :: "int \ (int \ int) set" where "int_ge_less_than d = {(z', z). d \ z' \ z' < z}" lemma wf_int_ge_less_than: "wf (int_ge_less_than d)" proof - have "int_ge_less_than d \ measure (\z. nat (z - d))" by (auto simp add: int_ge_less_than_def) then show ?thesis by (rule wf_subset [OF wf_measure]) qed text \ This variant looks odd, but is typical of the relations suggested by RankFinder.\ definition int_ge_less_than2 :: "int \ (int \ int) set" where "int_ge_less_than2 d = {(z',z). d \ z \ z' < z}" lemma wf_int_ge_less_than2: "wf (int_ge_less_than2 d)" proof - have "int_ge_less_than2 d \ measure (\z. nat (1 + z - d))" by (auto simp add: int_ge_less_than2_def) then show ?thesis by (rule wf_subset [OF wf_measure]) qed (* `set:int': dummy construction *) theorem int_ge_induct [case_names base step, induct set: int]: fixes i :: int assumes ge: "k \ i" and base: "P k" and step: "\i. k \ i \ P i \ P (i + 1)" shows "P i" proof - have "\i::int. n = nat (i - k) \ k \ i \ P i" for n proof (induct n) case 0 then have "i = k" by arith with base show "P i" by simp next case (Suc n) then have "n = nat ((i - 1) - k)" by arith moreover have k: "k \ i - 1" using Suc.prems by arith ultimately have "P (i - 1)" by (rule Suc.hyps) from step [OF k this] show ?case by simp qed with ge show ?thesis by fast qed (* `set:int': dummy construction *) theorem int_gr_induct [case_names base step, induct set: int]: fixes i k :: int assumes "k < i" "P (k + 1)" "\i. k < i \ P i \ P (i + 1)" shows "P i" proof - have "k+1 \ i" using assms by auto then show ?thesis by (induction i rule: int_ge_induct) (auto simp: assms) qed theorem int_le_induct [consumes 1, case_names base step]: fixes i k :: int assumes le: "i \ k" and base: "P k" and step: "\i. i \ k \ P i \ P (i - 1)" shows "P i" proof - have "\i::int. n = nat(k-i) \ i \ k \ P i" for n proof (induct n) case 0 then have "i = k" by arith with base show "P i" by simp next case (Suc n) then have "n = nat (k - (i + 1))" by arith moreover have k: "i + 1 \ k" using Suc.prems by arith ultimately have "P (i + 1)" by (rule Suc.hyps) from step[OF k this] show ?case by simp qed with le show ?thesis by fast qed theorem int_less_induct [consumes 1, case_names base step]: fixes i k :: int assumes "i < k" "P (k - 1)" "\i. i < k \ P i \ P (i - 1)" shows "P i" proof - have "i \ k-1" using assms by auto then show ?thesis by (induction i rule: int_le_induct) (auto simp: assms) qed theorem int_induct [case_names base step1 step2]: fixes k :: int assumes base: "P k" and step1: "\i. k \ i \ P i \ P (i + 1)" and step2: "\i. k \ i \ P i \ P (i - 1)" shows "P i" proof - have "i \ k \ i \ k" by arith then show ?thesis proof assume "i \ k" then show ?thesis using base by (rule int_ge_induct) (fact step1) next assume "i \ k" then show ?thesis using base by (rule int_le_induct) (fact step2) qed qed subsection \Intermediate value theorems\ lemma nat_ivt_aux: "\\if (Suc i) - f i\ \ 1; f 0 \ k; k \ f n\ \ \i \ n. f i = k" for m n :: nat and k :: int proof (induct n) case (Suc n) show ?case proof (cases "k = f (Suc n)") case False with Suc have "k \ f n" by auto with Suc show ?thesis by (auto simp add: abs_if split: if_split_asm intro: le_SucI) qed (use Suc in auto) qed auto lemma nat_intermed_int_val: fixes m n :: nat and k :: int assumes "\i. m \ i \ i < n \ \f (Suc i) - f i\ \ 1" "m \ n" "f m \ k" "k \ f n" shows "\i. m \ i \ i \ n \ f i = k" proof - obtain i where "i \ n - m" "k = f (m + i)" using nat_ivt_aux [of "n - m" "f \ plus m" k] assms by auto with assms show ?thesis by (rule_tac x = "m + i" in exI) auto qed lemma nat0_intermed_int_val: "\i\n. f i = k" if "\if (i + 1) - f i\ \ 1" "f 0 \ k" "k \ f n" for n :: nat and k :: int using nat_intermed_int_val [of 0 n f k] that by auto subsection \Products and 1, by T. M. Rasmussen\ lemma abs_zmult_eq_1: fixes m n :: int assumes mn: "\m * n\ = 1" shows "\m\ = 1" proof - from mn have 0: "m \ 0" "n \ 0" by auto have "\ 2 \ \m\" proof assume "2 \ \m\" then have "2 * \n\ \ \m\ * \n\" by (simp add: mult_mono 0) also have "\ = \m * n\" by (simp add: abs_mult) also from mn have "\ = 1" by simp finally have "2 * \n\ \ 1" . with 0 show "False" by arith qed with 0 show ?thesis by auto qed lemma pos_zmult_eq_1_iff_lemma: "m * n = 1 \ m = 1 \ m = - 1" for m n :: int using abs_zmult_eq_1 [of m n] by arith lemma pos_zmult_eq_1_iff: fixes m n :: int assumes "0 < m" shows "m * n = 1 \ m = 1 \ n = 1" proof - from assms have "m * n = 1 \ m = 1" by (auto dest: pos_zmult_eq_1_iff_lemma) then show ?thesis by (auto dest: pos_zmult_eq_1_iff_lemma) qed lemma zmult_eq_1_iff: "m * n = 1 \ (m = 1 \ n = 1) \ (m = - 1 \ n = - 1)" (is "?L = ?R") for m n :: int proof assume L: ?L show ?R using pos_zmult_eq_1_iff_lemma [OF L] L by force qed auto lemma infinite_UNIV_int [simp]: "\ finite (UNIV::int set)" proof assume "finite (UNIV::int set)" moreover have "inj (\i::int. 2 * i)" by (rule injI) simp ultimately have "surj (\i::int. 2 * i)" by (rule finite_UNIV_inj_surj) then obtain i :: int where "1 = 2 * i" by (rule surjE) then show False by (simp add: pos_zmult_eq_1_iff) qed subsection \The divides relation\ lemma zdvd_antisym_nonneg: "0 \ m \ 0 \ n \ m dvd n \ n dvd m \ m = n" for m n :: int by (auto simp add: dvd_def mult.assoc zero_le_mult_iff zmult_eq_1_iff) lemma zdvd_antisym_abs: fixes a b :: int assumes "a dvd b" and "b dvd a" shows "\a\ = \b\" proof (cases "a = 0") case True with assms show ?thesis by simp next case False from \a dvd b\ obtain k where k: "b = a * k" unfolding dvd_def by blast from \b dvd a\ obtain k' where k': "a = b * k'" unfolding dvd_def by blast from k k' have "a = a * k * k'" by simp with mult_cancel_left1[where c="a" and b="k*k'"] have kk': "k * k' = 1" using \a \ 0\ by (simp add: mult.assoc) then have "k = 1 \ k' = 1 \ k = -1 \ k' = -1" by (simp add: zmult_eq_1_iff) with k k' show ?thesis by auto qed lemma zdvd_zdiffD: "k dvd m - n \ k dvd n \ k dvd m" for k m n :: int using dvd_add_right_iff [of k "- n" m] by simp lemma zdvd_reduce: "k dvd n + k * m \ k dvd n" for k m n :: int using dvd_add_times_triv_right_iff [of k n m] by (simp add: ac_simps) lemma dvd_imp_le_int: fixes d i :: int assumes "i \ 0" and "d dvd i" shows "\d\ \ \i\" proof - from \d dvd i\ obtain k where "i = d * k" .. with \i \ 0\ have "k \ 0" by auto then have "1 \ \k\" and "0 \ \d\" by auto then have "\d\ * 1 \ \d\ * \k\" by (rule mult_left_mono) with \i = d * k\ show ?thesis by (simp add: abs_mult) qed lemma zdvd_not_zless: fixes m n :: int assumes "0 < m" and "m < n" shows "\ n dvd m" proof from assms have "0 < n" by auto assume "n dvd m" then obtain k where k: "m = n * k" .. with \0 < m\ have "0 < n * k" by auto with \0 < n\ have "0 < k" by (simp add: zero_less_mult_iff) with k \0 < n\ \m < n\ have "n * k < n * 1" by simp with \0 < n\ \0 < k\ show False unfolding mult_less_cancel_left by auto qed lemma zdvd_mult_cancel: fixes k m n :: int assumes d: "k * m dvd k * n" and "k \ 0" shows "m dvd n" proof - from d obtain h where h: "k * n = k * m * h" unfolding dvd_def by blast have "n = m * h" proof (rule ccontr) assume "\ ?thesis" with \k \ 0\ have "k * n \ k * (m * h)" by simp with h show False by (simp add: mult.assoc) qed then show ?thesis by simp qed lemma int_dvd_int_iff [simp]: "int m dvd int n \ m dvd n" proof - have "m dvd n" if "int n = int m * k" for k proof (cases k) case (nonneg q) with that have "n = m * q" by (simp del: of_nat_mult add: of_nat_mult [symmetric]) then show ?thesis .. next case (neg q) with that have "int n = int m * (- int (Suc q))" by simp also have "\ = - (int m * int (Suc q))" by (simp only: mult_minus_right) also have "\ = - int (m * Suc q)" by (simp only: of_nat_mult [symmetric]) finally have "- int (m * Suc q) = int n" .. then show ?thesis by (simp only: negative_eq_positive) auto qed then show ?thesis by (auto simp add: dvd_def) qed lemma dvd_nat_abs_iff [simp]: "n dvd nat \k\ \ int n dvd k" proof - have "n dvd nat \k\ \ int n dvd int (nat \k\)" by (simp only: int_dvd_int_iff) then show ?thesis by simp qed lemma nat_abs_dvd_iff [simp]: "nat \k\ dvd n \ k dvd int n" proof - have "nat \k\ dvd n \ int (nat \k\) dvd int n" by (simp only: int_dvd_int_iff) then show ?thesis by simp qed lemma zdvd1_eq [simp]: "x dvd 1 \ \x\ = 1" (is "?lhs \ ?rhs") for x :: int proof assume ?lhs then have "nat \x\ dvd nat \1\" by (simp only: nat_abs_dvd_iff) simp then have "nat \x\ = 1" by simp then show ?rhs by (cases "x < 0") simp_all next assume ?rhs then have "x = 1 \ x = - 1" by auto then show ?lhs by (auto intro: dvdI) qed lemma zdvd_mult_cancel1: fixes m :: int assumes mp: "m \ 0" shows "m * n dvd m \ \n\ = 1" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (cases "n > 0") (auto simp add: minus_equation_iff) next assume ?lhs then have "m * n dvd m * 1" by simp from zdvd_mult_cancel[OF this mp] show ?rhs by (simp only: zdvd1_eq) qed lemma nat_dvd_iff: "nat z dvd m \ (if 0 \ z then z dvd int m else m = 0)" using nat_abs_dvd_iff [of z m] by (cases "z \ 0") auto lemma eq_nat_nat_iff: "0 \ z \ 0 \ z' \ nat z = nat z' \ z = z'" by (auto elim: nonneg_int_cases) lemma nat_power_eq: "0 \ z \ nat (z ^ n) = nat z ^ n" by (induct n) (simp_all add: nat_mult_distrib) lemma numeral_power_eq_nat_cancel_iff [simp]: "numeral x ^ n = nat y \ numeral x ^ n = y" using nat_eq_iff2 by auto lemma nat_eq_numeral_power_cancel_iff [simp]: "nat y = numeral x ^ n \ y = numeral x ^ n" using numeral_power_eq_nat_cancel_iff[of x n y] by (metis (mono_tags)) lemma numeral_power_le_nat_cancel_iff [simp]: "numeral x ^ n \ nat a \ numeral x ^ n \ a" using nat_le_eq_zle[of "numeral x ^ n" a] by (auto simp: nat_power_eq) lemma nat_le_numeral_power_cancel_iff [simp]: "nat a \ numeral x ^ n \ a \ numeral x ^ n" by (simp add: nat_le_iff) lemma numeral_power_less_nat_cancel_iff [simp]: "numeral x ^ n < nat a \ numeral x ^ n < a" using nat_less_eq_zless[of "numeral x ^ n" a] by (auto simp: nat_power_eq) lemma nat_less_numeral_power_cancel_iff [simp]: "nat a < numeral x ^ n \ a < numeral x ^ n" using nat_less_eq_zless[of a "numeral x ^ n"] by (cases "a < 0") (auto simp: nat_power_eq less_le_trans[where y=0]) lemma zdvd_imp_le: "z \ n" if "z dvd n" "0 < n" for n z :: int proof (cases n) case (nonneg n) show ?thesis by (cases z) (use nonneg dvd_imp_le that in auto) qed (use that in auto) lemma zdvd_period: fixes a d :: int assumes "a dvd d" shows "a dvd (x + t) \ a dvd ((x + c * d) + t)" (is "?lhs \ ?rhs") proof - from assms have "a dvd (x + t) \ a dvd ((x + t) + c * d)" by (simp add: dvd_add_left_iff) then show ?thesis by (simp add: ac_simps) qed subsection \Powers with integer exponents\ text \ The following allows writing powers with an integer exponent. While the type signature is very generic, most theorems will assume that the underlying type is a division ring or a field. The notation `powi' is inspired by the `powr' notation for real/complex exponentiation. \ definition power_int :: "'a :: {inverse, power} \ int \ 'a" (infixr "powi" 80) where "power_int x n = (if n \ 0 then x ^ nat n else inverse x ^ (nat (-n)))" lemma power_int_0_right [simp]: "power_int x 0 = 1" and power_int_1_right [simp]: "power_int (y :: 'a :: {power, inverse, monoid_mult}) 1 = y" and power_int_minus1_right [simp]: "power_int (y :: 'a :: {power, inverse, monoid_mult}) (-1) = inverse y" by (simp_all add: power_int_def) lemma power_int_of_nat [simp]: "power_int x (int n) = x ^ n" by (simp add: power_int_def) lemma power_int_numeral [simp]: "power_int x (numeral n) = x ^ numeral n" by (simp add: power_int_def) lemma int_cases4 [case_names nonneg neg]: fixes m :: int obtains n where "m = int n" | n where "n > 0" "m = -int n" proof (cases "m \ 0") case True thus ?thesis using that(1)[of "nat m"] by auto next case False thus ?thesis using that(2)[of "nat (-m)"] by auto qed context assumes "SORT_CONSTRAINT('a::division_ring)" begin lemma power_int_minus: "power_int (x::'a) (-n) = inverse (power_int x n)" by (auto simp: power_int_def power_inverse) lemma power_int_eq_0_iff [simp]: "power_int (x::'a) n = 0 \ x = 0 \ n \ 0" by (auto simp: power_int_def) lemma power_int_0_left_If: "power_int (0 :: 'a) m = (if m = 0 then 1 else 0)" by (auto simp: power_int_def) lemma power_int_0_left [simp]: "m \ 0 \ power_int (0 :: 'a) m = 0" by (simp add: power_int_0_left_If) lemma power_int_1_left [simp]: "power_int 1 n = (1 :: 'a :: division_ring)" by (auto simp: power_int_def) lemma power_diff_conv_inverse: "x \ 0 \ m \ n \ (x :: 'a) ^ (n - m) = x ^ n * inverse x ^ m" by (simp add: field_simps flip: power_add) lemma power_mult_inverse_distrib: "x ^ m * inverse (x :: 'a) = inverse x * x ^ m" proof (cases "x = 0") case [simp]: False show ?thesis proof (cases m) case (Suc m') have "x ^ Suc m' * inverse x = x ^ m'" by (subst power_Suc2) (auto simp: mult.assoc) also have "\ = inverse x * x ^ Suc m'" by (subst power_Suc) (auto simp: mult.assoc [symmetric]) finally show ?thesis using Suc by simp qed auto qed auto lemma power_mult_power_inverse_commute: "x ^ m * inverse (x :: 'a) ^ n = inverse x ^ n * x ^ m" proof (induction n) case (Suc n) have "x ^ m * inverse x ^ Suc n = (x ^ m * inverse x ^ n) * inverse x" by (simp only: power_Suc2 mult.assoc) also have "x ^ m * inverse x ^ n = inverse x ^ n * x ^ m" by (rule Suc) also have "\ * inverse x = (inverse x ^ n * inverse x) * x ^ m" by (simp add: mult.assoc power_mult_inverse_distrib) also have "\ = inverse x ^ (Suc n) * x ^ m" by (simp only: power_Suc2) finally show ?case . qed auto lemma power_int_add: assumes "x \ 0 \ m + n \ 0" shows "power_int (x::'a) (m + n) = power_int x m * power_int x n" proof (cases "x = 0") case True thus ?thesis using assms by (auto simp: power_int_0_left_If) next case [simp]: False show ?thesis proof (cases m n rule: int_cases4[case_product int_cases4]) case (nonneg_nonneg a b) thus ?thesis by (auto simp: power_int_def nat_add_distrib power_add) next case (nonneg_neg a b) thus ?thesis by (auto simp: power_int_def nat_diff_distrib not_le power_diff_conv_inverse power_mult_power_inverse_commute) next case (neg_nonneg a b) thus ?thesis by (auto simp: power_int_def nat_diff_distrib not_le power_diff_conv_inverse power_mult_power_inverse_commute) next case (neg_neg a b) thus ?thesis by (auto simp: power_int_def nat_add_distrib add.commute simp flip: power_add) qed qed lemma power_int_add_1: assumes "x \ 0 \ m \ -1" shows "power_int (x::'a) (m + 1) = power_int x m * x" using assms by (subst power_int_add) auto lemma power_int_add_1': assumes "x \ 0 \ m \ -1" shows "power_int (x::'a) (m + 1) = x * power_int x m" using assms by (subst add.commute, subst power_int_add) auto lemma power_int_commutes: "power_int (x :: 'a) n * x = x * power_int x n" by (cases "x = 0") (auto simp flip: power_int_add_1 power_int_add_1') lemma power_int_inverse [field_simps, field_split_simps, divide_simps]: "power_int (inverse (x :: 'a)) n = inverse (power_int x n)" by (auto simp: power_int_def power_inverse) lemma power_int_mult: "power_int (x :: 'a) (m * n) = power_int (power_int x m) n" by (auto simp: power_int_def zero_le_mult_iff simp flip: power_mult power_inverse nat_mult_distrib) end context assumes "SORT_CONSTRAINT('a::field)" begin lemma power_int_diff: assumes "x \ 0 \ m \ n" shows "power_int (x::'a) (m - n) = power_int x m / power_int x n" using power_int_add[of x m "-n"] assms by (auto simp: field_simps power_int_minus) lemma power_int_minus_mult: "x \ 0 \ n \ 0 \ power_int (x :: 'a) (n - 1) * x = power_int x n" by (auto simp flip: power_int_add_1) lemma power_int_mult_distrib: "power_int (x * y :: 'a) m = power_int x m * power_int y m" by (auto simp: power_int_def power_mult_distrib) lemmas power_int_mult_distrib_numeral1 = power_int_mult_distrib [where x = "numeral w" for w, simp] lemmas power_int_mult_distrib_numeral2 = power_int_mult_distrib [where y = "numeral w" for w, simp] lemma power_int_divide_distrib: "power_int (x / y :: 'a) m = power_int x m / power_int y m" using power_int_mult_distrib[of x "inverse y" m] unfolding power_int_inverse by (simp add: field_simps) end lemma power_int_add_numeral [simp]: "power_int x (numeral m) * power_int x (numeral n) = power_int x (numeral (m + n))" for x :: "'a :: division_ring" by (simp add: power_int_add [symmetric]) lemma power_int_add_numeral2 [simp]: "power_int x (numeral m) * (power_int x (numeral n) * b) = power_int x (numeral (m + n)) * b" for x :: "'a :: division_ring" by (simp add: mult.assoc [symmetric]) lemma power_int_mult_numeral [simp]: "power_int (power_int x (numeral m)) (numeral n) = power_int x (numeral (m * n))" for x :: "'a :: division_ring" by (simp only: numeral_mult power_int_mult) lemma power_int_not_zero: "(x :: 'a :: division_ring) \ 0 \ n = 0 \ power_int x n \ 0" by (subst power_int_eq_0_iff) auto lemma power_int_one_over [field_simps, field_split_simps, divide_simps]: "power_int (1 / x :: 'a :: division_ring) n = 1 / power_int x n" using power_int_inverse[of x] by (simp add: divide_inverse) context assumes "SORT_CONSTRAINT('a :: linordered_field)" begin lemma power_int_numeral_neg_numeral [simp]: "power_int (numeral m) (-numeral n) = (inverse (numeral (Num.pow m n)) :: 'a)" by (simp add: power_int_minus) lemma zero_less_power_int [simp]: "0 < (x :: 'a) \ 0 < power_int x n" by (auto simp: power_int_def) lemma zero_le_power_int [simp]: "0 \ (x :: 'a) \ 0 \ power_int x n" by (auto simp: power_int_def) lemma power_int_mono: "(x :: 'a) \ y \ n \ 0 \ 0 \ x \ power_int x n \ power_int y n" by (cases n rule: int_cases4) (auto intro: power_mono) lemma one_le_power_int [simp]: "1 \ (x :: 'a) \ n \ 0 \ 1 \ power_int x n" using power_int_mono [of 1 x n] by simp lemma power_int_le_one: "0 \ (x :: 'a) \ n \ 0 \ x \ 1 \ power_int x n \ 1" using power_int_mono [of x 1 n] by simp lemma power_int_le_imp_le_exp: assumes gt1: "1 < (x :: 'a :: linordered_field)" assumes "power_int x m \ power_int x n" "n \ 0" shows "m \ n" proof (cases "m < 0") case True with \n \ 0\ show ?thesis by simp next case False with assms have "x ^ nat m \ x ^ nat n" by (simp add: power_int_def) from gt1 and this show ?thesis using False \n \ 0\ by auto qed lemma power_int_le_imp_less_exp: assumes gt1: "1 < (x :: 'a :: linordered_field)" assumes "power_int x m < power_int x n" "n \ 0" shows "m < n" proof (cases "m < 0") case True with \n \ 0\ show ?thesis by simp next case False with assms have "x ^ nat m < x ^ nat n" by (simp add: power_int_def) from gt1 and this show ?thesis using False \n \ 0\ by auto qed lemma power_int_strict_mono: "(a :: 'a :: linordered_field) < b \ 0 \ a \ 0 < n \ power_int a n < power_int b n" by (auto simp: power_int_def intro!: power_strict_mono) lemma power_int_mono_iff [simp]: fixes a b :: "'a :: linordered_field" shows "\a \ 0; b \ 0; n > 0\ \ power_int a n \ power_int b n \ a \ b" by (auto simp: power_int_def intro!: power_strict_mono) lemma power_int_strict_increasing: fixes a :: "'a :: linordered_field" assumes "n < N" "1 < a" shows "power_int a N > power_int a n" proof - have *: "a ^ nat (N - n) > a ^ 0" using assms by (intro power_strict_increasing) auto have "power_int a N = power_int a n * power_int a (N - n)" using assms by (simp flip: power_int_add) also have "\ > power_int a n * 1" using assms * by (intro mult_strict_left_mono zero_less_power_int) (auto simp: power_int_def) finally show ?thesis by simp qed lemma power_int_increasing: fixes a :: "'a :: linordered_field" assumes "n \ N" "a \ 1" shows "power_int a N \ power_int a n" proof - have *: "a ^ nat (N - n) \ a ^ 0" using assms by (intro power_increasing) auto have "power_int a N = power_int a n * power_int a (N - n)" using assms by (simp flip: power_int_add) also have "\ \ power_int a n * 1" using assms * by (intro mult_left_mono) (auto simp: power_int_def) finally show ?thesis by simp qed lemma power_int_strict_decreasing: fixes a :: "'a :: linordered_field" assumes "n < N" "0 < a" "a < 1" shows "power_int a N < power_int a n" proof - have *: "a ^ nat (N - n) < a ^ 0" using assms by (intro power_strict_decreasing) auto have "power_int a N = power_int a n * power_int a (N - n)" using assms by (simp flip: power_int_add) also have "\ < power_int a n * 1" using assms * by (intro mult_strict_left_mono zero_less_power_int) (auto simp: power_int_def) finally show ?thesis by simp qed lemma power_int_decreasing: fixes a :: "'a :: linordered_field" assumes "n \ N" "0 \ a" "a \ 1" "a \ 0 \ N \ 0 \ n = 0" shows "power_int a N \ power_int a n" proof (cases "a = 0") case False have *: "a ^ nat (N - n) \ a ^ 0" using assms by (intro power_decreasing) auto have "power_int a N = power_int a n * power_int a (N - n)" using assms False by (simp flip: power_int_add) also have "\ \ power_int a n * 1" using assms * by (intro mult_left_mono) (auto simp: power_int_def) finally show ?thesis by simp qed (use assms in \auto simp: power_int_0_left_If\) lemma one_less_power_int: "1 < (a :: 'a) \ 0 < n \ 1 < power_int a n" using power_int_strict_increasing[of 0 n a] by simp lemma power_int_abs: "\power_int a n :: 'a\ = power_int \a\ n" by (auto simp: power_int_def power_abs) lemma power_int_sgn [simp]: "sgn (power_int a n :: 'a) = power_int (sgn a) n" by (auto simp: power_int_def) lemma abs_power_int_minus [simp]: "\power_int (- a) n :: 'a\ = \power_int a n\" by (simp add: power_int_abs) lemma power_int_strict_antimono: assumes "(a :: 'a :: linordered_field) < b" "0 < a" "n < 0" shows "power_int a n > power_int b n" proof - have "inverse (power_int a (-n)) > inverse (power_int b (-n))" using assms by (intro less_imp_inverse_less power_int_strict_mono zero_less_power_int) auto thus ?thesis by (simp add: power_int_minus) qed lemma power_int_antimono: assumes "(a :: 'a :: linordered_field) \ b" "0 < a" "n < 0" shows "power_int a n \ power_int b n" using power_int_strict_antimono[of a b n] assms by (cases "a = b") auto end subsection \Finiteness of intervals\ lemma finite_interval_int1 [iff]: "finite {i :: int. a \ i \ i \ b}" proof (cases "a \ b") case True then show ?thesis proof (induct b rule: int_ge_induct) case base have "{i. a \ i \ i \ a} = {a}" by auto then show ?case by simp next case (step b) then have "{i. a \ i \ i \ b + 1} = {i. a \ i \ i \ b} \ {b + 1}" by auto with step show ?case by simp qed next case False then show ?thesis by (metis (lifting, no_types) Collect_empty_eq finite.emptyI order_trans) qed lemma finite_interval_int2 [iff]: "finite {i :: int. a \ i \ i < b}" by (rule rev_finite_subset[OF finite_interval_int1[of "a" "b"]]) auto lemma finite_interval_int3 [iff]: "finite {i :: int. a < i \ i \ b}" by (rule rev_finite_subset[OF finite_interval_int1[of "a" "b"]]) auto lemma finite_interval_int4 [iff]: "finite {i :: int. a < i \ i < b}" by (rule rev_finite_subset[OF finite_interval_int1[of "a" "b"]]) auto subsection \Configuration of the code generator\ text \Constructors\ definition Pos :: "num \ int" where [simp, code_abbrev]: "Pos = numeral" definition Neg :: "num \ int" where [simp, code_abbrev]: "Neg n = - (Pos n)" code_datatype "0::int" Pos Neg text \Auxiliary operations.\ definition dup :: "int \ int" where [simp]: "dup k = k + k" lemma dup_code [code]: "dup 0 = 0" "dup (Pos n) = Pos (Num.Bit0 n)" "dup (Neg n) = Neg (Num.Bit0 n)" by (simp_all add: numeral_Bit0) definition sub :: "num \ num \ int" where [simp]: "sub m n = numeral m - numeral n" lemma sub_code [code]: "sub Num.One Num.One = 0" "sub (Num.Bit0 m) Num.One = Pos (Num.BitM m)" "sub (Num.Bit1 m) Num.One = Pos (Num.Bit0 m)" "sub Num.One (Num.Bit0 n) = Neg (Num.BitM n)" "sub Num.One (Num.Bit1 n) = Neg (Num.Bit0 n)" "sub (Num.Bit0 m) (Num.Bit0 n) = dup (sub m n)" "sub (Num.Bit1 m) (Num.Bit1 n) = dup (sub m n)" "sub (Num.Bit1 m) (Num.Bit0 n) = dup (sub m n) + 1" "sub (Num.Bit0 m) (Num.Bit1 n) = dup (sub m n) - 1" by (simp_all only: sub_def dup_def numeral.simps Pos_def Neg_def numeral_BitM) lemma sub_BitM_One_eq: \(Num.sub (Num.BitM n) num.One) = 2 * (Num.sub n Num.One :: int)\ by (cases n) simp_all text \Implementations.\ lemma one_int_code [code]: "1 = Pos Num.One" by simp lemma plus_int_code [code]: "k + 0 = k" "0 + l = l" "Pos m + Pos n = Pos (m + n)" "Pos m + Neg n = sub m n" "Neg m + Pos n = sub n m" "Neg m + Neg n = Neg (m + n)" for k l :: int by simp_all lemma uminus_int_code [code]: "uminus 0 = (0::int)" "uminus (Pos m) = Neg m" "uminus (Neg m) = Pos m" by simp_all lemma minus_int_code [code]: "k - 0 = k" "0 - l = uminus l" "Pos m - Pos n = sub m n" "Pos m - Neg n = Pos (m + n)" "Neg m - Pos n = Neg (m + n)" "Neg m - Neg n = sub n m" for k l :: int by simp_all lemma times_int_code [code]: "k * 0 = 0" "0 * l = 0" "Pos m * Pos n = Pos (m * n)" "Pos m * Neg n = Neg (m * n)" "Neg m * Pos n = Neg (m * n)" "Neg m * Neg n = Pos (m * n)" for k l :: int by simp_all instantiation int :: equal begin definition "HOL.equal k l \ k = (l::int)" instance by standard (rule equal_int_def) end lemma equal_int_code [code]: "HOL.equal 0 (0::int) \ True" "HOL.equal 0 (Pos l) \ False" "HOL.equal 0 (Neg l) \ False" "HOL.equal (Pos k) 0 \ False" "HOL.equal (Pos k) (Pos l) \ HOL.equal k l" "HOL.equal (Pos k) (Neg l) \ False" "HOL.equal (Neg k) 0 \ False" "HOL.equal (Neg k) (Pos l) \ False" "HOL.equal (Neg k) (Neg l) \ HOL.equal k l" by (auto simp add: equal) lemma equal_int_refl [code nbe]: "HOL.equal k k \ True" for k :: int by (fact equal_refl) lemma less_eq_int_code [code]: "0 \ (0::int) \ True" "0 \ Pos l \ True" "0 \ Neg l \ False" "Pos k \ 0 \ False" "Pos k \ Pos l \ k \ l" "Pos k \ Neg l \ False" "Neg k \ 0 \ True" "Neg k \ Pos l \ True" "Neg k \ Neg l \ l \ k" by simp_all lemma less_int_code [code]: "0 < (0::int) \ False" "0 < Pos l \ True" "0 < Neg l \ False" "Pos k < 0 \ False" "Pos k < Pos l \ k < l" "Pos k < Neg l \ False" "Neg k < 0 \ True" "Neg k < Pos l \ True" "Neg k < Neg l \ l < k" by simp_all lemma nat_code [code]: "nat (Int.Neg k) = 0" "nat 0 = 0" "nat (Int.Pos k) = nat_of_num k" by (simp_all add: nat_of_num_numeral) lemma (in ring_1) of_int_code [code]: "of_int (Int.Neg k) = - numeral k" "of_int 0 = 0" "of_int (Int.Pos k) = numeral k" by simp_all text \Serializer setup.\ code_identifier code_module Int \ (SML) Arith and (OCaml) Arith and (Haskell) Arith quickcheck_params [default_type = int] hide_const (open) Pos Neg sub dup text \De-register \int\ as a quotient type:\ lifting_update int.lifting lifting_forget int.lifting subsection \Duplicates\ lemmas int_sum = of_nat_sum [where 'a=int] lemmas int_prod = of_nat_prod [where 'a=int] lemmas zle_int = of_nat_le_iff [where 'a=int] lemmas int_int_eq = of_nat_eq_iff [where 'a=int] lemmas nonneg_eq_int = nonneg_int_cases lemmas double_eq_0_iff = double_zero lemmas int_distrib = distrib_right [of z1 z2 w] distrib_left [of w z1 z2] left_diff_distrib [of z1 z2 w] right_diff_distrib [of w z1 z2] for z1 z2 w :: int end diff --git a/src/HOL/Library/Numeral_Type.thy b/src/HOL/Library/Numeral_Type.thy --- a/src/HOL/Library/Numeral_Type.thy +++ b/src/HOL/Library/Numeral_Type.thy @@ -1,579 +1,604 @@ (* Title: HOL/Library/Numeral_Type.thy Author: Brian Huffman *) section \Numeral Syntax for Types\ theory Numeral_Type imports Cardinality begin subsection \Numeral Types\ typedef num0 = "UNIV :: nat set" .. typedef num1 = "UNIV :: unit set" .. typedef 'a bit0 = "{0 ..< 2 * int CARD('a::finite)}" proof show "0 \ {0 ..< 2 * int CARD('a)}" by simp qed typedef 'a bit1 = "{0 ..< 1 + 2 * int CARD('a::finite)}" proof show "0 \ {0 ..< 1 + 2 * int CARD('a)}" by simp qed lemma card_num0 [simp]: "CARD (num0) = 0" unfolding type_definition.card [OF type_definition_num0] by simp lemma infinite_num0: "\ finite (UNIV :: num0 set)" using card_num0[unfolded card_eq_0_iff] by simp lemma card_num1 [simp]: "CARD(num1) = 1" unfolding type_definition.card [OF type_definition_num1] by (simp only: card_UNIV_unit) lemma card_bit0 [simp]: "CARD('a bit0) = 2 * CARD('a::finite)" unfolding type_definition.card [OF type_definition_bit0] by simp lemma card_bit1 [simp]: "CARD('a bit1) = Suc (2 * CARD('a::finite))" unfolding type_definition.card [OF type_definition_bit1] by simp subsection \@{typ num1}\ instance num1 :: finite proof show "finite (UNIV::num1 set)" unfolding type_definition.univ [OF type_definition_num1] using finite by (rule finite_imageI) qed instantiation num1 :: CARD_1 begin instance proof show "CARD(num1) = 1" by auto qed end lemma num1_eq_iff: "(x::num1) = (y::num1) \ True" by (induct x, induct y) simp instantiation num1 :: "{comm_ring,comm_monoid_mult,numeral}" begin instance by standard (simp_all add: num1_eq_iff) end lemma num1_eqI: fixes a::num1 shows "a = b" by(simp add: num1_eq_iff) lemma num1_eq1 [simp]: fixes a::num1 shows "a = 1" by (rule num1_eqI) lemma forall_1[simp]: "(\i::num1. P i) \ P 1" by (metis (full_types) num1_eq_iff) lemma ex_1[simp]: "(\x::num1. P x) \ P 1" by auto (metis (full_types) num1_eq_iff) instantiation num1 :: linorder begin definition "a < b \ Rep_num1 a < Rep_num1 b" definition "a \ b \ Rep_num1 a \ Rep_num1 b" instance by intro_classes (auto simp: less_eq_num1_def less_num1_def intro: num1_eqI) end instance num1 :: wellorder by intro_classes (auto simp: less_eq_num1_def less_num1_def) instance bit0 :: (finite) card2 proof show "finite (UNIV::'a bit0 set)" unfolding type_definition.univ [OF type_definition_bit0] by simp show "2 \ CARD('a bit0)" by simp qed instance bit1 :: (finite) card2 proof show "finite (UNIV::'a bit1 set)" unfolding type_definition.univ [OF type_definition_bit1] by simp show "2 \ CARD('a bit1)" by simp qed subsection \Locales for for modular arithmetic subtypes\ locale mod_type = fixes n :: int and Rep :: "'a::{zero,one,plus,times,uminus,minus} \ int" and Abs :: "int \ 'a::{zero,one,plus,times,uminus,minus}" assumes type: "type_definition Rep Abs {0.. n" by (rule Rep_less_n [THEN order_less_imp_le]) lemma Rep_inject_sym: "x = y \ Rep x = Rep y" by (rule type_definition.Rep_inject [OF type, symmetric]) lemma Rep_inverse: "Abs (Rep x) = x" by (rule type_definition.Rep_inverse [OF type]) lemma Abs_inverse: "m \ {0.. Rep (Abs m) = m" by (rule type_definition.Abs_inverse [OF type]) lemma Rep_Abs_mod: "Rep (Abs (m mod n)) = m mod n" by (simp add: Abs_inverse pos_mod_conj [OF size0]) lemma Rep_Abs_0: "Rep (Abs 0) = 0" by (simp add: Abs_inverse size0) lemma Rep_0: "Rep 0 = 0" by (simp add: zero_def Rep_Abs_0) lemma Rep_Abs_1: "Rep (Abs 1) = 1" by (simp add: Abs_inverse size1) lemma Rep_1: "Rep 1 = 1" by (simp add: one_def Rep_Abs_1) lemma Rep_mod: "Rep x mod n = Rep x" apply (rule_tac x=x in type_definition.Abs_cases [OF type]) apply (simp add: type_definition.Abs_inverse [OF type]) done lemmas Rep_simps = Rep_inject_sym Rep_inverse Rep_Abs_mod Rep_mod Rep_Abs_0 Rep_Abs_1 lemma comm_ring_1: "OFCLASS('a, comm_ring_1_class)" apply (intro_classes, unfold definitions) apply (simp_all add: Rep_simps mod_simps field_simps) done end locale mod_ring = mod_type n Rep Abs for n :: int and Rep :: "'a::{comm_ring_1} \ int" and Abs :: "int \ 'a::{comm_ring_1}" begin lemma of_nat_eq: "of_nat k = Abs (int k mod n)" apply (induct k) apply (simp add: zero_def) apply (simp add: Rep_simps add_def one_def mod_simps ac_simps) done lemma of_int_eq: "of_int z = Abs (z mod n)" apply (cases z rule: int_diff_cases) apply (simp add: Rep_simps of_nat_eq diff_def mod_simps) done lemma Rep_numeral: "Rep (numeral w) = numeral w mod n" using of_int_eq [of "numeral w"] by (simp add: Rep_inject_sym Rep_Abs_mod) lemma iszero_numeral: "iszero (numeral w::'a) \ numeral w mod n = 0" by (simp add: Rep_inject_sym Rep_numeral Rep_0 iszero_def) lemma cases: assumes 1: "\z. \(x::'a) = of_int z; 0 \ z; z < n\ \ P" shows "P" apply (cases x rule: type_definition.Abs_cases [OF type]) apply (rule_tac z="y" in 1) apply (simp_all add: of_int_eq) done lemma induct: "(\z. \0 \ z; z < n\ \ P (of_int z)) \ P (x::'a)" by (cases x rule: cases) simp +lemma UNIV_eq: "(UNIV :: 'a set) = Abs ` {0..Ring class instances\ text \ Unfortunately \ring_1\ instance is not possible for \<^typ>\num1\, since 0 and 1 are not distinct. \ instantiation bit0 and bit1 :: (finite) "{zero,one,plus,times,uminus,minus}" begin definition Abs_bit0' :: "int \ 'a bit0" where "Abs_bit0' x = Abs_bit0 (x mod int CARD('a bit0))" definition Abs_bit1' :: "int \ 'a bit1" where "Abs_bit1' x = Abs_bit1 (x mod int CARD('a bit1))" definition "0 = Abs_bit0 0" definition "1 = Abs_bit0 1" definition "x + y = Abs_bit0' (Rep_bit0 x + Rep_bit0 y)" definition "x * y = Abs_bit0' (Rep_bit0 x * Rep_bit0 y)" definition "x - y = Abs_bit0' (Rep_bit0 x - Rep_bit0 y)" definition "- x = Abs_bit0' (- Rep_bit0 x)" definition "0 = Abs_bit1 0" definition "1 = Abs_bit1 1" definition "x + y = Abs_bit1' (Rep_bit1 x + Rep_bit1 y)" definition "x * y = Abs_bit1' (Rep_bit1 x * Rep_bit1 y)" definition "x - y = Abs_bit1' (Rep_bit1 x - Rep_bit1 y)" definition "- x = Abs_bit1' (- Rep_bit1 x)" instance .. end interpretation bit0: mod_type "int CARD('a::finite bit0)" "Rep_bit0 :: 'a::finite bit0 \ int" "Abs_bit0 :: int \ 'a::finite bit0" apply (rule mod_type.intro) apply (simp add: type_definition_bit0) apply (rule one_less_int_card) apply (rule zero_bit0_def) apply (rule one_bit0_def) apply (rule plus_bit0_def [unfolded Abs_bit0'_def]) apply (rule times_bit0_def [unfolded Abs_bit0'_def]) apply (rule minus_bit0_def [unfolded Abs_bit0'_def]) apply (rule uminus_bit0_def [unfolded Abs_bit0'_def]) done interpretation bit1: mod_type "int CARD('a::finite bit1)" "Rep_bit1 :: 'a::finite bit1 \ int" "Abs_bit1 :: int \ 'a::finite bit1" apply (rule mod_type.intro) apply (simp add: type_definition_bit1) apply (rule one_less_int_card) apply (rule zero_bit1_def) apply (rule one_bit1_def) apply (rule plus_bit1_def [unfolded Abs_bit1'_def]) apply (rule times_bit1_def [unfolded Abs_bit1'_def]) apply (rule minus_bit1_def [unfolded Abs_bit1'_def]) apply (rule uminus_bit1_def [unfolded Abs_bit1'_def]) done instance bit0 :: (finite) comm_ring_1 by (rule bit0.comm_ring_1) instance bit1 :: (finite) comm_ring_1 by (rule bit1.comm_ring_1) interpretation bit0: mod_ring "int CARD('a::finite bit0)" "Rep_bit0 :: 'a::finite bit0 \ int" "Abs_bit0 :: int \ 'a::finite bit0" .. interpretation bit1: mod_ring "int CARD('a::finite bit1)" "Rep_bit1 :: 'a::finite bit1 \ int" "Abs_bit1 :: int \ 'a::finite bit1" .. text \Set up cases, induction, and arithmetic\ lemmas bit0_cases [case_names of_int, cases type: bit0] = bit0.cases lemmas bit1_cases [case_names of_int, cases type: bit1] = bit1.cases lemmas bit0_induct [case_names of_int, induct type: bit0] = bit0.induct lemmas bit1_induct [case_names of_int, induct type: bit1] = bit1.induct lemmas bit0_iszero_numeral [simp] = bit0.iszero_numeral lemmas bit1_iszero_numeral [simp] = bit1.iszero_numeral lemmas [simp] = eq_numeral_iff_iszero [where 'a="'a bit0"] for dummy :: "'a::finite" lemmas [simp] = eq_numeral_iff_iszero [where 'a="'a bit1"] for dummy :: "'a::finite" subsection \Order instances\ instantiation bit0 and bit1 :: (finite) linorder begin definition "a < b \ Rep_bit0 a < Rep_bit0 b" definition "a \ b \ Rep_bit0 a \ Rep_bit0 b" definition "a < b \ Rep_bit1 a < Rep_bit1 b" definition "a \ b \ Rep_bit1 a \ Rep_bit1 b" instance by(intro_classes) (auto simp add: less_eq_bit0_def less_bit0_def less_eq_bit1_def less_bit1_def Rep_bit0_inject Rep_bit1_inject) end lemma (in preorder) tranclp_less: "(<) \<^sup>+\<^sup>+ = (<)" by(auto simp add: fun_eq_iff intro: less_trans elim: tranclp.induct) instance bit0 and bit1 :: (finite) wellorder proof - have "wf {(x :: 'a bit0, y). x < y}" by(auto simp add: trancl_def tranclp_less intro!: finite_acyclic_wf acyclicI) thus "OFCLASS('a bit0, wellorder_class)" by(rule wf_wellorderI) intro_classes next have "wf {(x :: 'a bit1, y). x < y}" by(auto simp add: trancl_def tranclp_less intro!: finite_acyclic_wf acyclicI) thus "OFCLASS('a bit1, wellorder_class)" by(rule wf_wellorderI) intro_classes qed subsection \Code setup and type classes for code generation\ text \Code setup for \<^typ>\num0\ and \<^typ>\num1\\ definition Num0 :: num0 where "Num0 = Abs_num0 0" code_datatype Num0 instantiation num0 :: equal begin definition equal_num0 :: "num0 \ num0 \ bool" where "equal_num0 = (=)" instance by intro_classes (simp add: equal_num0_def) end lemma equal_num0_code [code]: "equal_class.equal Num0 Num0 = True" by(rule equal_refl) code_datatype "1 :: num1" instantiation num1 :: equal begin definition equal_num1 :: "num1 \ num1 \ bool" where "equal_num1 = (=)" instance by intro_classes (simp add: equal_num1_def) end lemma equal_num1_code [code]: "equal_class.equal (1 :: num1) 1 = True" by(rule equal_refl) instantiation num1 :: enum begin definition "enum_class.enum = [1 :: num1]" definition "enum_class.enum_all P = P (1 :: num1)" definition "enum_class.enum_ex P = P (1 :: num1)" instance by intro_classes (auto simp add: enum_num1_def enum_all_num1_def enum_ex_num1_def num1_eq_iff Ball_def) end instantiation num0 and num1 :: card_UNIV begin definition "finite_UNIV = Phantom(num0) False" definition "card_UNIV = Phantom(num0) 0" definition "finite_UNIV = Phantom(num1) True" definition "card_UNIV = Phantom(num1) 1" instance by intro_classes (simp_all add: finite_UNIV_num0_def card_UNIV_num0_def infinite_num0 finite_UNIV_num1_def card_UNIV_num1_def) end text \Code setup for \<^typ>\'a bit0\ and \<^typ>\'a bit1\\ declare bit0.Rep_inverse[code abstype] bit0.Rep_0[code abstract] bit0.Rep_1[code abstract] lemma Abs_bit0'_code [code abstract]: "Rep_bit0 (Abs_bit0' x :: 'a :: finite bit0) = x mod int (CARD('a bit0))" by(auto simp add: Abs_bit0'_def intro!: Abs_bit0_inverse) lemma inj_on_Abs_bit0: "inj_on (Abs_bit0 :: int \ 'a bit0) {0..<2 * int CARD('a :: finite)}" by(auto intro: inj_onI simp add: Abs_bit0_inject) declare bit1.Rep_inverse[code abstype] bit1.Rep_0[code abstract] bit1.Rep_1[code abstract] lemma Abs_bit1'_code [code abstract]: "Rep_bit1 (Abs_bit1' x :: 'a :: finite bit1) = x mod int (CARD('a bit1))" by(auto simp add: Abs_bit1'_def intro!: Abs_bit1_inverse) lemma inj_on_Abs_bit1: "inj_on (Abs_bit1 :: int \ 'a bit1) {0..<1 + 2 * int CARD('a :: finite)}" by(auto intro: inj_onI simp add: Abs_bit1_inject) instantiation bit0 and bit1 :: (finite) equal begin definition "equal_class.equal x y \ Rep_bit0 x = Rep_bit0 y" definition "equal_class.equal x y \ Rep_bit1 x = Rep_bit1 y" instance by intro_classes (simp_all add: equal_bit0_def equal_bit1_def Rep_bit0_inject Rep_bit1_inject) end instantiation bit0 :: (finite) enum begin definition "(enum_class.enum :: 'a bit0 list) = map (Abs_bit0' \ int) (upt 0 (CARD('a bit0)))" definition "enum_class.enum_all P = (\b :: 'a bit0 \ set enum_class.enum. P b)" definition "enum_class.enum_ex P = (\b :: 'a bit0 \ set enum_class.enum. P b)" instance proof show "distinct (enum_class.enum :: 'a bit0 list)" by (simp add: enum_bit0_def distinct_map inj_on_def Abs_bit0'_def Abs_bit0_inject) let ?Abs = "Abs_bit0 :: _ \ 'a bit0" interpret type_definition Rep_bit0 ?Abs "{0..<2 * int CARD('a)}" by (fact type_definition_bit0) have "UNIV = ?Abs ` {0..<2 * int CARD('a)}" by (simp add: Abs_image) also have "\ = ?Abs ` (int ` {0..<2 * CARD('a)})" by (simp add: image_int_atLeastLessThan) also have "\ = (?Abs \ int) ` {0..<2 * CARD('a)}" by (simp add: image_image cong: image_cong) also have "\ = set enum_class.enum" by (simp add: enum_bit0_def Abs_bit0'_def cong: image_cong_simp) finally show univ_eq: "(UNIV :: 'a bit0 set) = set enum_class.enum" . fix P :: "'a bit0 \ bool" show "enum_class.enum_all P = Ball UNIV P" and "enum_class.enum_ex P = Bex UNIV P" by(simp_all add: enum_all_bit0_def enum_ex_bit0_def univ_eq) qed end instantiation bit1 :: (finite) enum begin definition "(enum_class.enum :: 'a bit1 list) = map (Abs_bit1' \ int) (upt 0 (CARD('a bit1)))" definition "enum_class.enum_all P = (\b :: 'a bit1 \ set enum_class.enum. P b)" definition "enum_class.enum_ex P = (\b :: 'a bit1 \ set enum_class.enum. P b)" instance proof(intro_classes) show "distinct (enum_class.enum :: 'a bit1 list)" by(simp only: Abs_bit1'_def zmod_int[symmetric] enum_bit1_def distinct_map Suc_eq_plus1 card_bit1 o_apply inj_on_def) (clarsimp simp add: Abs_bit1_inject) let ?Abs = "Abs_bit1 :: _ \ 'a bit1" interpret type_definition Rep_bit1 ?Abs "{0..<1 + 2 * int CARD('a)}" by (fact type_definition_bit1) have "UNIV = ?Abs ` {0..<1 + 2 * int CARD('a)}" by (simp add: Abs_image) also have "\ = ?Abs ` (int ` {0..<1 + 2 * CARD('a)})" by (simp add: image_int_atLeastLessThan) also have "\ = (?Abs \ int) ` {0..<1 + 2 * CARD('a)}" by (simp add: image_image cong: image_cong) finally show univ_eq: "(UNIV :: 'a bit1 set) = set enum_class.enum" by (simp only: enum_bit1_def set_map set_upt) (simp add: Abs_bit1'_def cong: image_cong_simp) fix P :: "'a bit1 \ bool" show "enum_class.enum_all P = Ball UNIV P" and "enum_class.enum_ex P = Bex UNIV P" by(simp_all add: enum_all_bit1_def enum_ex_bit1_def univ_eq) qed end instantiation bit0 and bit1 :: (finite) finite_UNIV begin definition "finite_UNIV = Phantom('a bit0) True" definition "finite_UNIV = Phantom('a bit1) True" instance by intro_classes (simp_all add: finite_UNIV_bit0_def finite_UNIV_bit1_def) end instantiation bit0 and bit1 :: ("{finite,card_UNIV}") card_UNIV begin definition "card_UNIV = Phantom('a bit0) (2 * of_phantom (card_UNIV :: 'a card_UNIV))" definition "card_UNIV = Phantom('a bit1) (1 + 2 * of_phantom (card_UNIV :: 'a card_UNIV))" instance by intro_classes (simp_all add: card_UNIV_bit0_def card_UNIV_bit1_def card_UNIV) end subsection \Syntax\ syntax "_NumeralType" :: "num_token => type" ("_") "_NumeralType0" :: type ("0") "_NumeralType1" :: type ("1") translations (type) "1" == (type) "num1" (type) "0" == (type) "num0" parse_translation \ let fun mk_bintype n = let fun mk_bit 0 = Syntax.const \<^type_syntax>\bit0\ | mk_bit 1 = Syntax.const \<^type_syntax>\bit1\; fun bin_of n = if n = 1 then Syntax.const \<^type_syntax>\num1\ else if n = 0 then Syntax.const \<^type_syntax>\num0\ else if n = ~1 then raise TERM ("negative type numeral", []) else let val (q, r) = Integer.div_mod n 2; in mk_bit r $ bin_of q end; in bin_of n end; fun numeral_tr [Free (str, _)] = mk_bintype (the (Int.fromString str)) | numeral_tr ts = raise TERM ("numeral_tr", ts); in [(\<^syntax_const>\_NumeralType\, K numeral_tr)] end \ print_translation \ let fun int_of [] = 0 | int_of (b :: bs) = b + 2 * int_of bs; fun bin_of (Const (\<^type_syntax>\num0\, _)) = [] | bin_of (Const (\<^type_syntax>\num1\, _)) = [1] | bin_of (Const (\<^type_syntax>\bit0\, _) $ bs) = 0 :: bin_of bs | bin_of (Const (\<^type_syntax>\bit1\, _) $ bs) = 1 :: bin_of bs | bin_of t = raise TERM ("bin_of", [t]); fun bit_tr' b [t] = let val rev_digs = b :: bin_of t handle TERM _ => raise Match val i = int_of rev_digs; val num = string_of_int (abs i); in Syntax.const \<^syntax_const>\_NumeralType\ $ Syntax.free num end | bit_tr' b _ = raise Match; in [(\<^type_syntax>\bit0\, K (bit_tr' 0)), (\<^type_syntax>\bit1\, K (bit_tr' 1))] end \ subsection \Examples\ lemma "CARD(0) = 0" by simp lemma "CARD(17) = 17" by simp +lemma "CHAR(23) = 23" by simp lemma "8 * 11 ^ 3 - 6 = (2::5)" by simp end diff --git a/src/HOL/Rat.thy b/src/HOL/Rat.thy --- a/src/HOL/Rat.thy +++ b/src/HOL/Rat.thy @@ -1,1153 +1,1167 @@ (* Title: HOL/Rat.thy Author: Markus Wenzel, TU Muenchen *) section \Rational numbers\ theory Rat imports Archimedean_Field begin subsection \Rational numbers as quotient\ subsubsection \Construction of the type of rational numbers\ definition ratrel :: "(int \ int) \ (int \ int) \ bool" where "ratrel = (\x y. snd x \ 0 \ snd y \ 0 \ fst x * snd y = fst y * snd x)" lemma ratrel_iff [simp]: "ratrel x y \ snd x \ 0 \ snd y \ 0 \ fst x * snd y = fst y * snd x" by (simp add: ratrel_def) lemma exists_ratrel_refl: "\x. ratrel x x" by (auto intro!: one_neq_zero) lemma symp_ratrel: "symp ratrel" by (simp add: ratrel_def symp_def) lemma transp_ratrel: "transp ratrel" proof (rule transpI, unfold split_paired_all) fix a b a' b' a'' b'' :: int assume *: "ratrel (a, b) (a', b')" assume **: "ratrel (a', b') (a'', b'')" have "b' * (a * b'') = b'' * (a * b')" by simp also from * have "a * b' = a' * b" by auto also have "b'' * (a' * b) = b * (a' * b'')" by simp also from ** have "a' * b'' = a'' * b'" by auto also have "b * (a'' * b') = b' * (a'' * b)" by simp finally have "b' * (a * b'') = b' * (a'' * b)" . moreover from ** have "b' \ 0" by auto ultimately have "a * b'' = a'' * b" by simp with * ** show "ratrel (a, b) (a'', b'')" by auto qed lemma part_equivp_ratrel: "part_equivp ratrel" by (rule part_equivpI [OF exists_ratrel_refl symp_ratrel transp_ratrel]) quotient_type rat = "int \ int" / partial: "ratrel" morphisms Rep_Rat Abs_Rat by (rule part_equivp_ratrel) lemma Domainp_cr_rat [transfer_domain_rule]: "Domainp pcr_rat = (\x. snd x \ 0)" by (simp add: rat.domain_eq) subsubsection \Representation and basic operations\ lift_definition Fract :: "int \ int \ rat" is "\a b. if b = 0 then (0, 1) else (a, b)" by simp lemma eq_rat: "\a b c d. b \ 0 \ d \ 0 \ Fract a b = Fract c d \ a * d = c * b" "\a. Fract a 0 = Fract 0 1" "\a c. Fract 0 a = Fract 0 c" by (transfer, simp)+ lemma Rat_cases [case_names Fract, cases type: rat]: assumes that: "\a b. q = Fract a b \ b > 0 \ coprime a b \ C" shows C proof - obtain a b :: int where q: "q = Fract a b" and b: "b \ 0" by transfer simp let ?a = "a div gcd a b" let ?b = "b div gcd a b" from b have "?b * gcd a b = b" by simp with b have "?b \ 0" by fastforce with q b have q2: "q = Fract ?a ?b" by (simp add: eq_rat dvd_div_mult mult.commute [of a]) from b have coprime: "coprime ?a ?b" by (auto intro: div_gcd_coprime) show C proof (cases "b > 0") case True then have "?b > 0" by (simp add: nonneg1_imp_zdiv_pos_iff) from q2 this coprime show C by (rule that) next case False have "q = Fract (- ?a) (- ?b)" unfolding q2 by transfer simp moreover from False b have "- ?b > 0" by (simp add: pos_imp_zdiv_neg_iff) moreover from coprime have "coprime (- ?a) (- ?b)" by simp ultimately show C by (rule that) qed qed lemma Rat_induct [case_names Fract, induct type: rat]: assumes "\a b. b > 0 \ coprime a b \ P (Fract a b)" shows "P q" using assms by (cases q) simp instantiation rat :: field begin lift_definition zero_rat :: "rat" is "(0, 1)" by simp lift_definition one_rat :: "rat" is "(1, 1)" by simp lemma Zero_rat_def: "0 = Fract 0 1" by transfer simp lemma One_rat_def: "1 = Fract 1 1" by transfer simp lift_definition plus_rat :: "rat \ rat \ rat" is "\x y. (fst x * snd y + fst y * snd x, snd x * snd y)" by (auto simp: distrib_right) (simp add: ac_simps) lemma add_rat [simp]: assumes "b \ 0" and "d \ 0" shows "Fract a b + Fract c d = Fract (a * d + c * b) (b * d)" using assms by transfer simp lift_definition uminus_rat :: "rat \ rat" is "\x. (- fst x, snd x)" by simp lemma minus_rat [simp]: "- Fract a b = Fract (- a) b" by transfer simp lemma minus_rat_cancel [simp]: "Fract (- a) (- b) = Fract a b" by (cases "b = 0") (simp_all add: eq_rat) definition diff_rat_def: "q - r = q + - r" for q r :: rat lemma diff_rat [simp]: "b \ 0 \ d \ 0 \ Fract a b - Fract c d = Fract (a * d - c * b) (b * d)" by (simp add: diff_rat_def) lift_definition times_rat :: "rat \ rat \ rat" is "\x y. (fst x * fst y, snd x * snd y)" by (simp add: ac_simps) lemma mult_rat [simp]: "Fract a b * Fract c d = Fract (a * c) (b * d)" by transfer simp lemma mult_rat_cancel: "c \ 0 \ Fract (c * a) (c * b) = Fract a b" by transfer simp lift_definition inverse_rat :: "rat \ rat" is "\x. if fst x = 0 then (0, 1) else (snd x, fst x)" by (auto simp add: mult.commute) lemma inverse_rat [simp]: "inverse (Fract a b) = Fract b a" by transfer simp definition divide_rat_def: "q div r = q * inverse r" for q r :: rat lemma divide_rat [simp]: "Fract a b div Fract c d = Fract (a * d) (b * c)" by (simp add: divide_rat_def) instance proof fix q r s :: rat show "(q * r) * s = q * (r * s)" by transfer simp show "q * r = r * q" by transfer simp show "1 * q = q" by transfer simp show "(q + r) + s = q + (r + s)" by transfer (simp add: algebra_simps) show "q + r = r + q" by transfer simp show "0 + q = q" by transfer simp show "- q + q = 0" by transfer simp show "q - r = q + - r" by (fact diff_rat_def) show "(q + r) * s = q * s + r * s" by transfer (simp add: algebra_simps) show "(0::rat) \ 1" by transfer simp show "inverse q * q = 1" if "q \ 0" using that by transfer simp show "q div r = q * inverse r" by (fact divide_rat_def) show "inverse 0 = (0::rat)" by transfer simp qed end (* We cannot state these two rules earlier because of pending sort hypotheses *) lemma div_add_self1_no_field [simp]: assumes "NO_MATCH (x :: 'b :: field) b" "(b :: 'a :: euclidean_semiring_cancel) \ 0" shows "(b + a) div b = a div b + 1" using assms(2) by (fact div_add_self1) lemma div_add_self2_no_field [simp]: assumes "NO_MATCH (x :: 'b :: field) b" "(b :: 'a :: euclidean_semiring_cancel) \ 0" shows "(a + b) div b = a div b + 1" using assms(2) by (fact div_add_self2) lemma of_nat_rat: "of_nat k = Fract (of_nat k) 1" by (induct k) (simp_all add: Zero_rat_def One_rat_def) lemma of_int_rat: "of_int k = Fract k 1" by (cases k rule: int_diff_cases) (simp add: of_nat_rat) lemma Fract_of_nat_eq: "Fract (of_nat k) 1 = of_nat k" by (rule of_nat_rat [symmetric]) lemma Fract_of_int_eq: "Fract k 1 = of_int k" by (rule of_int_rat [symmetric]) lemma rat_number_collapse: "Fract 0 k = 0" "Fract 1 1 = 1" "Fract (numeral w) 1 = numeral w" "Fract (- numeral w) 1 = - numeral w" "Fract (- 1) 1 = - 1" "Fract k 0 = 0" using Fract_of_int_eq [of "numeral w"] and Fract_of_int_eq [of "- numeral w"] by (simp_all add: Zero_rat_def One_rat_def eq_rat) lemma rat_number_expand: "0 = Fract 0 1" "1 = Fract 1 1" "numeral k = Fract (numeral k) 1" "- 1 = Fract (- 1) 1" "- numeral k = Fract (- numeral k) 1" by (simp_all add: rat_number_collapse) lemma Rat_cases_nonzero [case_names Fract 0]: assumes Fract: "\a b. q = Fract a b \ b > 0 \ a \ 0 \ coprime a b \ C" and 0: "q = 0 \ C" shows C proof (cases "q = 0") case True then show C using 0 by auto next case False then obtain a b where *: "q = Fract a b" "b > 0" "coprime a b" by (cases q) auto with False have "0 \ Fract a b" by simp with \b > 0\ have "a \ 0" by (simp add: Zero_rat_def eq_rat) with Fract * show C by blast qed subsubsection \Function \normalize\\ lemma Fract_coprime: "Fract (a div gcd a b) (b div gcd a b) = Fract a b" proof (cases "b = 0") case True then show ?thesis by (simp add: eq_rat) next case False moreover have "b div gcd a b * gcd a b = b" by (rule dvd_div_mult_self) simp ultimately have "b div gcd a b * gcd a b \ 0" by simp then have "b div gcd a b \ 0" by fastforce with False show ?thesis by (simp add: eq_rat dvd_div_mult mult.commute [of a]) qed definition normalize :: "int \ int \ int \ int" where "normalize p = (if snd p > 0 then (let a = gcd (fst p) (snd p) in (fst p div a, snd p div a)) else if snd p = 0 then (0, 1) else (let a = - gcd (fst p) (snd p) in (fst p div a, snd p div a)))" lemma normalize_crossproduct: assumes "q \ 0" "s \ 0" assumes "normalize (p, q) = normalize (r, s)" shows "p * s = r * q" proof - have *: "p * s = q * r" if "p * gcd r s = sgn (q * s) * r * gcd p q" and "q * gcd r s = sgn (q * s) * s * gcd p q" proof - from that have "(p * gcd r s) * (sgn (q * s) * s * gcd p q) = (q * gcd r s) * (sgn (q * s) * r * gcd p q)" by simp with assms show ?thesis by (auto simp add: ac_simps sgn_mult sgn_0_0) qed from assms show ?thesis by (auto simp: normalize_def Let_def dvd_div_div_eq_mult mult.commute sgn_mult split: if_splits intro: *) qed lemma normalize_eq: "normalize (a, b) = (p, q) \ Fract p q = Fract a b" by (auto simp: normalize_def Let_def Fract_coprime dvd_div_neg rat_number_collapse split: if_split_asm) lemma normalize_denom_pos: "normalize r = (p, q) \ q > 0" by (auto simp: normalize_def Let_def dvd_div_neg pos_imp_zdiv_neg_iff nonneg1_imp_zdiv_pos_iff split: if_split_asm) lemma normalize_coprime: "normalize r = (p, q) \ coprime p q" by (auto simp: normalize_def Let_def dvd_div_neg div_gcd_coprime split: if_split_asm) lemma normalize_stable [simp]: "q > 0 \ coprime p q \ normalize (p, q) = (p, q)" by (simp add: normalize_def) lemma normalize_denom_zero [simp]: "normalize (p, 0) = (0, 1)" by (simp add: normalize_def) lemma normalize_negative [simp]: "q < 0 \ normalize (p, q) = normalize (- p, - q)" by (simp add: normalize_def Let_def dvd_div_neg dvd_neg_div) text\ Decompose a fraction into normalized, i.e. coprime numerator and denominator: \ definition quotient_of :: "rat \ int \ int" where "quotient_of x = (THE pair. x = Fract (fst pair) (snd pair) \ snd pair > 0 \ coprime (fst pair) (snd pair))" lemma quotient_of_unique: "\!p. r = Fract (fst p) (snd p) \ snd p > 0 \ coprime (fst p) (snd p)" proof (cases r) case (Fract a b) then have "r = Fract (fst (a, b)) (snd (a, b)) \ snd (a, b) > 0 \ coprime (fst (a, b)) (snd (a, b))" by auto then show ?thesis proof (rule ex1I) fix p assume r: "r = Fract (fst p) (snd p) \ snd p > 0 \ coprime (fst p) (snd p)" obtain c d where p: "p = (c, d)" by (cases p) with r have Fract': "r = Fract c d" "d > 0" "coprime c d" by simp_all have "(c, d) = (a, b)" proof (cases "a = 0") case True with Fract Fract' show ?thesis by (simp add: eq_rat) next case False with Fract Fract' have *: "c * b = a * d" and "c \ 0" by (auto simp add: eq_rat) then have "c * b > 0 \ a * d > 0" by auto with \b > 0\ \d > 0\ have "a > 0 \ c > 0" by (simp add: zero_less_mult_iff) with \a \ 0\ \c \ 0\ have sgn: "sgn a = sgn c" by (auto simp add: not_less) from \coprime a b\ \coprime c d\ have "\a\ * \d\ = \c\ * \b\ \ \a\ = \c\ \ \d\ = \b\" by (simp add: coprime_crossproduct_int) with \b > 0\ \d > 0\ have "\a\ * d = \c\ * b \ \a\ = \c\ \ d = b" by simp then have "a * sgn a * d = c * sgn c * b \ a * sgn a = c * sgn c \ d = b" by (simp add: abs_sgn) with sgn * show ?thesis by (auto simp add: sgn_0_0) qed with p show "p = (a, b)" by simp qed qed lemma quotient_of_Fract [code]: "quotient_of (Fract a b) = normalize (a, b)" proof - have "Fract a b = Fract (fst (normalize (a, b))) (snd (normalize (a, b)))" (is ?Fract) by (rule sym) (auto intro: normalize_eq) moreover have "0 < snd (normalize (a, b))" (is ?denom_pos) by (cases "normalize (a, b)") (rule normalize_denom_pos, simp) moreover have "coprime (fst (normalize (a, b))) (snd (normalize (a, b)))" (is ?coprime) by (rule normalize_coprime) simp ultimately have "?Fract \ ?denom_pos \ ?coprime" by blast then have "(THE p. Fract a b = Fract (fst p) (snd p) \ 0 < snd p \ coprime (fst p) (snd p)) = normalize (a, b)" by (rule the1_equality [OF quotient_of_unique]) then show ?thesis by (simp add: quotient_of_def) qed lemma quotient_of_number [simp]: "quotient_of 0 = (0, 1)" "quotient_of 1 = (1, 1)" "quotient_of (numeral k) = (numeral k, 1)" "quotient_of (- 1) = (- 1, 1)" "quotient_of (- numeral k) = (- numeral k, 1)" by (simp_all add: rat_number_expand quotient_of_Fract) lemma quotient_of_eq: "quotient_of (Fract a b) = (p, q) \ Fract p q = Fract a b" by (simp add: quotient_of_Fract normalize_eq) lemma quotient_of_denom_pos: "quotient_of r = (p, q) \ q > 0" by (cases r) (simp add: quotient_of_Fract normalize_denom_pos) lemma quotient_of_denom_pos': "snd (quotient_of r) > 0" using quotient_of_denom_pos [of r] by (simp add: prod_eq_iff) lemma quotient_of_coprime: "quotient_of r = (p, q) \ coprime p q" by (cases r) (simp add: quotient_of_Fract normalize_coprime) lemma quotient_of_inject: assumes "quotient_of a = quotient_of b" shows "a = b" proof - obtain p q r s where a: "a = Fract p q" and b: "b = Fract r s" and "q > 0" and "s > 0" by (cases a, cases b) with assms show ?thesis by (simp add: eq_rat quotient_of_Fract normalize_crossproduct) qed lemma quotient_of_inject_eq: "quotient_of a = quotient_of b \ a = b" by (auto simp add: quotient_of_inject) subsubsection \Various\ lemma Fract_of_int_quotient: "Fract k l = of_int k / of_int l" by (simp add: Fract_of_int_eq [symmetric]) lemma Fract_add_one: "n \ 0 \ Fract (m + n) n = Fract m n + 1" by (simp add: rat_number_expand) lemma quotient_of_div: assumes r: "quotient_of r = (n,d)" shows "r = of_int n / of_int d" proof - from theI'[OF quotient_of_unique[of r], unfolded r[unfolded quotient_of_def]] have "r = Fract n d" by simp then show ?thesis using Fract_of_int_quotient by simp qed subsubsection \The ordered field of rational numbers\ lift_definition positive :: "rat \ bool" is "\x. 0 < fst x * snd x" proof clarsimp fix a b c d :: int assume "b \ 0" and "d \ 0" and "a * d = c * b" then have "a * d * b * d = c * b * b * d" by simp then have "a * b * d\<^sup>2 = c * d * b\<^sup>2" unfolding power2_eq_square by (simp add: ac_simps) then have "0 < a * b * d\<^sup>2 \ 0 < c * d * b\<^sup>2" by simp then show "0 < a * b \ 0 < c * d" using \b \ 0\ and \d \ 0\ by (simp add: zero_less_mult_iff) qed lemma positive_zero: "\ positive 0" by transfer simp lemma positive_add: "positive x \ positive y \ positive (x + y)" apply transfer apply (auto simp add: zero_less_mult_iff add_pos_pos add_neg_neg mult_pos_neg mult_neg_pos mult_neg_neg) done lemma positive_mult: "positive x \ positive y \ positive (x * y)" apply transfer by (metis fst_conv mult.commute mult_pos_neg2 snd_conv zero_less_mult_iff) lemma positive_minus: "\ positive x \ x \ 0 \ positive (- x)" by transfer (auto simp: neq_iff zero_less_mult_iff mult_less_0_iff) instantiation rat :: linordered_field begin definition "x < y \ positive (y - x)" definition "x \ y \ x < y \ x = y" for x y :: rat definition "\a\ = (if a < 0 then - a else a)" for a :: rat definition "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)" for a :: rat instance proof fix a b c :: rat show "\a\ = (if a < 0 then - a else a)" by (rule abs_rat_def) show "a < b \ a \ b \ \ b \ a" unfolding less_eq_rat_def less_rat_def using positive_add positive_zero by force show "a \ a" unfolding less_eq_rat_def by simp show "a \ b \ b \ c \ a \ c" unfolding less_eq_rat_def less_rat_def using positive_add by fastforce show "a \ b \ b \ a \ a = b" unfolding less_eq_rat_def less_rat_def using positive_add positive_zero by fastforce show "a \ b \ c + a \ c + b" unfolding less_eq_rat_def less_rat_def by auto show "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)" by (rule sgn_rat_def) show "a \ b \ b \ a" unfolding less_eq_rat_def less_rat_def by (auto dest!: positive_minus) show "a < b \ 0 < c \ c * a < c * b" unfolding less_rat_def by (metis diff_zero positive_mult right_diff_distrib') qed end instantiation rat :: distrib_lattice begin definition "(inf :: rat \ rat \ rat) = min" definition "(sup :: rat \ rat \ rat) = max" instance by standard (auto simp add: inf_rat_def sup_rat_def max_min_distrib2) end lemma positive_rat: "positive (Fract a b) \ 0 < a * b" by transfer simp lemma less_rat [simp]: "b \ 0 \ d \ 0 \ Fract a b < Fract c d \ (a * d) * (b * d) < (c * b) * (b * d)" by (simp add: less_rat_def positive_rat algebra_simps) lemma le_rat [simp]: "b \ 0 \ d \ 0 \ Fract a b \ Fract c d \ (a * d) * (b * d) \ (c * b) * (b * d)" by (simp add: le_less eq_rat) lemma abs_rat [simp, code]: "\Fract a b\ = Fract \a\ \b\" by (auto simp add: abs_rat_def zabs_def Zero_rat_def not_less le_less eq_rat zero_less_mult_iff) lemma sgn_rat [simp, code]: "sgn (Fract a b) = of_int (sgn a * sgn b)" unfolding Fract_of_int_eq by (auto simp: zsgn_def sgn_rat_def Zero_rat_def eq_rat) (auto simp: rat_number_collapse not_less le_less zero_less_mult_iff) lemma Rat_induct_pos [case_names Fract, induct type: rat]: assumes step: "\a b. 0 < b \ P (Fract a b)" shows "P q" proof (cases q) case (Fract a b) have step': "P (Fract a b)" if b: "b < 0" for a b :: int proof - from b have "0 < - b" by simp then have "P (Fract (- a) (- b))" by (rule step) then show "P (Fract a b)" by (simp add: order_less_imp_not_eq [OF b]) qed from Fract show "P q" by (auto simp add: linorder_neq_iff step step') qed lemma zero_less_Fract_iff: "0 < b \ 0 < Fract a b \ 0 < a" by (simp add: Zero_rat_def zero_less_mult_iff) lemma Fract_less_zero_iff: "0 < b \ Fract a b < 0 \ a < 0" by (simp add: Zero_rat_def mult_less_0_iff) lemma zero_le_Fract_iff: "0 < b \ 0 \ Fract a b \ 0 \ a" by (simp add: Zero_rat_def zero_le_mult_iff) lemma Fract_le_zero_iff: "0 < b \ Fract a b \ 0 \ a \ 0" by (simp add: Zero_rat_def mult_le_0_iff) lemma one_less_Fract_iff: "0 < b \ 1 < Fract a b \ b < a" by (simp add: One_rat_def mult_less_cancel_right_disj) lemma Fract_less_one_iff: "0 < b \ Fract a b < 1 \ a < b" by (simp add: One_rat_def mult_less_cancel_right_disj) lemma one_le_Fract_iff: "0 < b \ 1 \ Fract a b \ b \ a" by (simp add: One_rat_def mult_le_cancel_right) lemma Fract_le_one_iff: "0 < b \ Fract a b \ 1 \ a \ b" by (simp add: One_rat_def mult_le_cancel_right) subsubsection \Rationals are an Archimedean field\ lemma rat_floor_lemma: "of_int (a div b) \ Fract a b \ Fract a b < of_int (a div b + 1)" proof - have "Fract a b = of_int (a div b) + Fract (a mod b) b" by (cases "b = 0") (simp, simp add: of_int_rat) moreover have "0 \ Fract (a mod b) b \ Fract (a mod b) b < 1" unfolding Fract_of_int_quotient by (rule linorder_cases [of b 0]) (simp_all add: divide_nonpos_neg) ultimately show ?thesis by simp qed instance rat :: archimedean_field proof show "\z. r \ of_int z" for r :: rat proof (induct r) case (Fract a b) have "Fract a b \ of_int (a div b + 1)" using rat_floor_lemma [of a b] by simp then show "\z. Fract a b \ of_int z" .. qed qed instantiation rat :: floor_ceiling begin definition floor_rat :: "rat \ int" where"\x\ = (THE z. of_int z \ x \ x < of_int (z + 1))" for x :: rat instance proof show "of_int \x\ \ x \ x < of_int (\x\ + 1)" for x :: rat unfolding floor_rat_def using floor_exists1 by (rule theI') qed end lemma floor_Fract [simp]: "\Fract a b\ = a div b" by (simp add: Fract_of_int_quotient floor_divide_of_int_eq) subsection \Linear arithmetic setup\ declaration \ K (Lin_Arith.add_inj_thms @{thms of_int_le_iff [THEN iffD2] of_int_eq_iff [THEN iffD2]} (* not needed because x < (y::int) can be rewritten as x + 1 <= y: of_int_less_iff RS iffD2 *) #> Lin_Arith.add_inj_const (\<^const_name>\of_nat\, \<^typ>\nat \ rat\) #> Lin_Arith.add_inj_const (\<^const_name>\of_int\, \<^typ>\int \ rat\)) \ subsection \Embedding from Rationals to other Fields\ context field_char_0 begin lift_definition of_rat :: "rat \ 'a" is "\x. of_int (fst x) / of_int (snd x)" by (auto simp: nonzero_divide_eq_eq nonzero_eq_divide_eq) (simp only: of_int_mult [symmetric]) end lemma of_rat_rat: "b \ 0 \ of_rat (Fract a b) = of_int a / of_int b" by transfer simp lemma of_rat_0 [simp]: "of_rat 0 = 0" by transfer simp lemma of_rat_1 [simp]: "of_rat 1 = 1" by transfer simp lemma of_rat_add: "of_rat (a + b) = of_rat a + of_rat b" by transfer (simp add: add_frac_eq) lemma of_rat_minus: "of_rat (- a) = - of_rat a" by transfer simp lemma of_rat_neg_one [simp]: "of_rat (- 1) = - 1" by (simp add: of_rat_minus) lemma of_rat_diff: "of_rat (a - b) = of_rat a - of_rat b" using of_rat_add [of a "- b"] by (simp add: of_rat_minus) lemma of_rat_mult: "of_rat (a * b) = of_rat a * of_rat b" by transfer (simp add: divide_inverse nonzero_inverse_mult_distrib ac_simps) lemma of_rat_sum: "of_rat (\a\A. f a) = (\a\A. of_rat (f a))" by (induct rule: infinite_finite_induct) (auto simp: of_rat_add) lemma of_rat_prod: "of_rat (\a\A. f a) = (\a\A. of_rat (f a))" by (induct rule: infinite_finite_induct) (auto simp: of_rat_mult) lemma nonzero_of_rat_inverse: "a \ 0 \ of_rat (inverse a) = inverse (of_rat a)" by (rule inverse_unique [symmetric]) (simp add: of_rat_mult [symmetric]) lemma of_rat_inverse: "(of_rat (inverse a) :: 'a::field_char_0) = inverse (of_rat a)" by (cases "a = 0") (simp_all add: nonzero_of_rat_inverse) lemma nonzero_of_rat_divide: "b \ 0 \ of_rat (a / b) = of_rat a / of_rat b" by (simp add: divide_inverse of_rat_mult nonzero_of_rat_inverse) lemma of_rat_divide: "(of_rat (a / b) :: 'a::field_char_0) = of_rat a / of_rat b" by (cases "b = 0") (simp_all add: nonzero_of_rat_divide) lemma of_rat_power: "(of_rat (a ^ n) :: 'a::field_char_0) = of_rat a ^ n" by (induct n) (simp_all add: of_rat_mult) lemma of_rat_eq_iff [simp]: "of_rat a = of_rat b \ a = b" apply transfer apply (simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq flip: of_int_mult) done lemma of_rat_eq_0_iff [simp]: "of_rat a = 0 \ a = 0" using of_rat_eq_iff [of _ 0] by simp lemma zero_eq_of_rat_iff [simp]: "0 = of_rat a \ 0 = a" by simp lemma of_rat_eq_1_iff [simp]: "of_rat a = 1 \ a = 1" using of_rat_eq_iff [of _ 1] by simp lemma one_eq_of_rat_iff [simp]: "1 = of_rat a \ 1 = a" by simp lemma of_rat_less: "(of_rat r :: 'a::linordered_field) < of_rat s \ r < s" proof (induct r, induct s) fix a b c d :: int assume not_zero: "b > 0" "d > 0" then have "b * d > 0" by simp have of_int_divide_less_eq: "(of_int a :: 'a) / of_int b < of_int c / of_int d \ (of_int a :: 'a) * of_int d < of_int c * of_int b" using not_zero by (simp add: pos_less_divide_eq pos_divide_less_eq) show "(of_rat (Fract a b) :: 'a::linordered_field) < of_rat (Fract c d) \ Fract a b < Fract c d" using not_zero \b * d > 0\ by (simp add: of_rat_rat of_int_divide_less_eq of_int_mult [symmetric] del: of_int_mult) qed lemma of_rat_less_eq: "(of_rat r :: 'a::linordered_field) \ of_rat s \ r \ s" unfolding le_less by (auto simp add: of_rat_less) lemma of_rat_le_0_iff [simp]: "(of_rat r :: 'a::linordered_field) \ 0 \ r \ 0" using of_rat_less_eq [of r 0, where 'a = 'a] by simp lemma zero_le_of_rat_iff [simp]: "0 \ (of_rat r :: 'a::linordered_field) \ 0 \ r" using of_rat_less_eq [of 0 r, where 'a = 'a] by simp lemma of_rat_le_1_iff [simp]: "(of_rat r :: 'a::linordered_field) \ 1 \ r \ 1" using of_rat_less_eq [of r 1] by simp lemma one_le_of_rat_iff [simp]: "1 \ (of_rat r :: 'a::linordered_field) \ 1 \ r" using of_rat_less_eq [of 1 r] by simp lemma of_rat_less_0_iff [simp]: "(of_rat r :: 'a::linordered_field) < 0 \ r < 0" using of_rat_less [of r 0, where 'a = 'a] by simp lemma zero_less_of_rat_iff [simp]: "0 < (of_rat r :: 'a::linordered_field) \ 0 < r" using of_rat_less [of 0 r, where 'a = 'a] by simp lemma of_rat_less_1_iff [simp]: "(of_rat r :: 'a::linordered_field) < 1 \ r < 1" using of_rat_less [of r 1] by simp lemma one_less_of_rat_iff [simp]: "1 < (of_rat r :: 'a::linordered_field) \ 1 < r" using of_rat_less [of 1 r] by simp lemma of_rat_eq_id [simp]: "of_rat = id" proof show "of_rat a = id a" for a by (induct a) (simp add: of_rat_rat Fract_of_int_eq [symmetric]) qed lemma abs_of_rat [simp]: "\of_rat r\ = (of_rat \r\ :: 'a :: linordered_field)" by (cases "r \ 0") (simp_all add: not_le of_rat_minus) text \Collapse nested embeddings.\ lemma of_rat_of_nat_eq [simp]: "of_rat (of_nat n) = of_nat n" by (induct n) (simp_all add: of_rat_add) lemma of_rat_of_int_eq [simp]: "of_rat (of_int z) = of_int z" by (cases z rule: int_diff_cases) (simp add: of_rat_diff) lemma of_rat_numeral_eq [simp]: "of_rat (numeral w) = numeral w" using of_rat_of_int_eq [of "numeral w"] by simp lemma of_rat_neg_numeral_eq [simp]: "of_rat (- numeral w) = - numeral w" using of_rat_of_int_eq [of "- numeral w"] by simp lemma of_rat_floor [simp]: "\of_rat r\ = \r\" by (cases r) (simp add: Fract_of_int_quotient of_rat_divide floor_divide_of_int_eq) lemma of_rat_ceiling [simp]: "\of_rat r\ = \r\" using of_rat_floor [of "- r"] by (simp add: of_rat_minus ceiling_def) lemmas zero_rat = Zero_rat_def lemmas one_rat = One_rat_def abbreviation rat_of_nat :: "nat \ rat" where "rat_of_nat \ of_nat" abbreviation rat_of_int :: "int \ rat" where "rat_of_int \ of_int" subsection \The Set of Rational Numbers\ context field_char_0 begin definition Rats :: "'a set" ("\") where "\ = range of_rat" end lemma Rats_cases [cases set: Rats]: assumes "q \ \" obtains (of_rat) r where "q = of_rat r" proof - from \q \ \\ have "q \ range of_rat" by (simp only: Rats_def) then obtain r where "q = of_rat r" .. then show thesis .. qed +lemma Rats_cases': + assumes "(x :: 'a :: field_char_0) \ \" + obtains a b where "b > 0" "coprime a b" "x = of_int a / of_int b" +proof - + from assms obtain r where "x = of_rat r" + by (auto simp: Rats_def) + obtain a b where quot: "quotient_of r = (a,b)" by force + have "b > 0" using quotient_of_denom_pos[OF quot] . + moreover have "coprime a b" using quotient_of_coprime[OF quot] . + moreover have "x = of_int a / of_int b" unfolding \x = of_rat r\ + quotient_of_div[OF quot] by (simp add: of_rat_divide) + ultimately show ?thesis using that by blast +qed + lemma Rats_of_rat [simp]: "of_rat r \ \" by (simp add: Rats_def) lemma Rats_of_int [simp]: "of_int z \ \" by (subst of_rat_of_int_eq [symmetric]) (rule Rats_of_rat) lemma Ints_subset_Rats: "\ \ \" using Ints_cases Rats_of_int by blast lemma Rats_of_nat [simp]: "of_nat n \ \" by (subst of_rat_of_nat_eq [symmetric]) (rule Rats_of_rat) lemma Nats_subset_Rats: "\ \ \" using Ints_subset_Rats Nats_subset_Ints by blast lemma Rats_number_of [simp]: "numeral w \ \" by (subst of_rat_numeral_eq [symmetric]) (rule Rats_of_rat) lemma Rats_0 [simp]: "0 \ \" unfolding Rats_def by (rule range_eqI) (rule of_rat_0 [symmetric]) lemma Rats_1 [simp]: "1 \ \" unfolding Rats_def by (rule range_eqI) (rule of_rat_1 [symmetric]) lemma Rats_add [simp]: "a \ \ \ b \ \ \ a + b \ \" by (metis Rats_cases Rats_of_rat of_rat_add) lemma Rats_minus_iff [simp]: "- a \ \ \ a \ \" by (metis Rats_cases Rats_of_rat add.inverse_inverse of_rat_minus) lemma Rats_diff [simp]: "a \ \ \ b \ \ \ a - b \ \" by (metis Rats_add Rats_minus_iff diff_conv_add_uminus) lemma Rats_mult [simp]: "a \ \ \ b \ \ \ a * b \ \" by (metis Rats_cases Rats_of_rat of_rat_mult) lemma Rats_inverse [simp]: "a \ \ \ inverse a \ \" for a :: "'a::field_char_0" by (metis Rats_cases Rats_of_rat of_rat_inverse) lemma Rats_divide [simp]: "a \ \ \ b \ \ \ a / b \ \" for a b :: "'a::field_char_0" by (simp add: divide_inverse) lemma Rats_power [simp]: "a \ \ \ a ^ n \ \" for a :: "'a::field_char_0" by (metis Rats_cases Rats_of_rat of_rat_power) lemma Rats_induct [case_names of_rat, induct set: Rats]: "q \ \ \ (\r. P (of_rat r)) \ P q" by (rule Rats_cases) auto lemma Rats_infinite: "\ finite \" by (auto dest!: finite_imageD simp: inj_on_def infinite_UNIV_char_0 Rats_def) subsection \Implementation of rational numbers as pairs of integers\ text \Formal constructor\ definition Frct :: "int \ int \ rat" where [simp]: "Frct p = Fract (fst p) (snd p)" lemma [code abstype]: "Frct (quotient_of q) = q" by (cases q) (auto intro: quotient_of_eq) text \Numerals\ declare quotient_of_Fract [code abstract] definition of_int :: "int \ rat" where [code_abbrev]: "of_int = Int.of_int" hide_const (open) of_int lemma quotient_of_int [code abstract]: "quotient_of (Rat.of_int a) = (a, 1)" by (simp add: of_int_def of_int_rat quotient_of_Fract) lemma [code_unfold]: "numeral k = Rat.of_int (numeral k)" by (simp add: Rat.of_int_def) lemma [code_unfold]: "- numeral k = Rat.of_int (- numeral k)" by (simp add: Rat.of_int_def) lemma Frct_code_post [code_post]: "Frct (0, a) = 0" "Frct (a, 0) = 0" "Frct (1, 1) = 1" "Frct (numeral k, 1) = numeral k" "Frct (1, numeral k) = 1 / numeral k" "Frct (numeral k, numeral l) = numeral k / numeral l" "Frct (- a, b) = - Frct (a, b)" "Frct (a, - b) = - Frct (a, b)" "- (- Frct q) = Frct q" by (simp_all add: Fract_of_int_quotient) text \Operations\ lemma rat_zero_code [code abstract]: "quotient_of 0 = (0, 1)" by (simp add: Zero_rat_def quotient_of_Fract normalize_def) lemma rat_one_code [code abstract]: "quotient_of 1 = (1, 1)" by (simp add: One_rat_def quotient_of_Fract normalize_def) lemma rat_plus_code [code abstract]: "quotient_of (p + q) = (let (a, c) = quotient_of p; (b, d) = quotient_of q in normalize (a * d + b * c, c * d))" by (cases p, cases q) (simp add: quotient_of_Fract) lemma rat_uminus_code [code abstract]: "quotient_of (- p) = (let (a, b) = quotient_of p in (- a, b))" by (cases p) (simp add: quotient_of_Fract) lemma rat_minus_code [code abstract]: "quotient_of (p - q) = (let (a, c) = quotient_of p; (b, d) = quotient_of q in normalize (a * d - b * c, c * d))" by (cases p, cases q) (simp add: quotient_of_Fract) lemma rat_times_code [code abstract]: "quotient_of (p * q) = (let (a, c) = quotient_of p; (b, d) = quotient_of q in normalize (a * b, c * d))" by (cases p, cases q) (simp add: quotient_of_Fract) lemma rat_inverse_code [code abstract]: "quotient_of (inverse p) = (let (a, b) = quotient_of p in if a = 0 then (0, 1) else (sgn a * b, \a\))" proof (cases p) case (Fract a b) then show ?thesis by (cases "0::int" a rule: linorder_cases) (simp_all add: quotient_of_Fract ac_simps) qed lemma rat_divide_code [code abstract]: "quotient_of (p / q) = (let (a, c) = quotient_of p; (b, d) = quotient_of q in normalize (a * d, c * b))" by (cases p, cases q) (simp add: quotient_of_Fract) lemma rat_abs_code [code abstract]: "quotient_of \p\ = (let (a, b) = quotient_of p in (\a\, b))" by (cases p) (simp add: quotient_of_Fract) lemma rat_sgn_code [code abstract]: "quotient_of (sgn p) = (sgn (fst (quotient_of p)), 1)" proof (cases p) case (Fract a b) then show ?thesis by (cases "0::int" a rule: linorder_cases) (simp_all add: quotient_of_Fract) qed lemma rat_floor_code [code]: "\p\ = (let (a, b) = quotient_of p in a div b)" by (cases p) (simp add: quotient_of_Fract floor_Fract) instantiation rat :: equal begin definition [code]: "HOL.equal a b \ quotient_of a = quotient_of b" instance by standard (simp add: equal_rat_def quotient_of_inject_eq) lemma rat_eq_refl [code nbe]: "HOL.equal (r::rat) r \ True" by (rule equal_refl) end lemma rat_less_eq_code [code]: "p \ q \ (let (a, c) = quotient_of p; (b, d) = quotient_of q in a * d \ c * b)" by (cases p, cases q) (simp add: quotient_of_Fract mult.commute) lemma rat_less_code [code]: "p < q \ (let (a, c) = quotient_of p; (b, d) = quotient_of q in a * d < c * b)" by (cases p, cases q) (simp add: quotient_of_Fract mult.commute) lemma [code]: "of_rat p = (let (a, b) = quotient_of p in of_int a / of_int b)" by (cases p) (simp add: quotient_of_Fract of_rat_rat) text \Quickcheck\ context includes term_syntax begin definition valterm_fract :: "int \ (unit \ Code_Evaluation.term) \ int \ (unit \ Code_Evaluation.term) \ rat \ (unit \ Code_Evaluation.term)" where [code_unfold]: "valterm_fract k l = Code_Evaluation.valtermify Fract {\} k {\} l" end instantiation rat :: random begin context includes state_combinator_syntax begin definition "Quickcheck_Random.random i = Quickcheck_Random.random i \\ (\num. Random.range i \\ (\denom. Pair (let j = int_of_integer (integer_of_natural (denom + 1)) in valterm_fract num (j, \u. Code_Evaluation.term_of j))))" instance .. end end instantiation rat :: exhaustive begin definition "exhaustive_rat f d = Quickcheck_Exhaustive.exhaustive (\l. Quickcheck_Exhaustive.exhaustive (\k. f (Fract k (int_of_integer (integer_of_natural l) + 1))) d) d" instance .. end instantiation rat :: full_exhaustive begin definition "full_exhaustive_rat f d = Quickcheck_Exhaustive.full_exhaustive (\(l, _). Quickcheck_Exhaustive.full_exhaustive (\k. f (let j = int_of_integer (integer_of_natural l) + 1 in valterm_fract k (j, \_. Code_Evaluation.term_of j))) d) d" instance .. end instance rat :: partial_term_of .. lemma [code]: "partial_term_of (ty :: rat itself) (Quickcheck_Narrowing.Narrowing_variable p tt) \ Code_Evaluation.Free (STR ''_'') (Typerep.Typerep (STR ''Rat.rat'') [])" "partial_term_of (ty :: rat itself) (Quickcheck_Narrowing.Narrowing_constructor 0 [l, k]) \ Code_Evaluation.App (Code_Evaluation.Const (STR ''Rat.Frct'') (Typerep.Typerep (STR ''fun'') [Typerep.Typerep (STR ''Product_Type.prod'') [Typerep.Typerep (STR ''Int.int'') [], Typerep.Typerep (STR ''Int.int'') []], Typerep.Typerep (STR ''Rat.rat'') []])) (Code_Evaluation.App (Code_Evaluation.App (Code_Evaluation.Const (STR ''Product_Type.Pair'') (Typerep.Typerep (STR ''fun'') [Typerep.Typerep (STR ''Int.int'') [], Typerep.Typerep (STR ''fun'') [Typerep.Typerep (STR ''Int.int'') [], Typerep.Typerep (STR ''Product_Type.prod'') [Typerep.Typerep (STR ''Int.int'') [], Typerep.Typerep (STR ''Int.int'') []]]])) (partial_term_of (TYPE(int)) l)) (partial_term_of (TYPE(int)) k))" by (rule partial_term_of_anything)+ instantiation rat :: narrowing begin definition "narrowing = Quickcheck_Narrowing.apply (Quickcheck_Narrowing.apply (Quickcheck_Narrowing.cons (\nom denom. Fract nom denom)) narrowing) narrowing" instance .. end subsection \Setup for Nitpick\ declaration \ Nitpick_HOL.register_frac_type \<^type_name>\rat\ [(\<^const_name>\Abs_Rat\, \<^const_name>\Nitpick.Abs_Frac\), (\<^const_name>\zero_rat_inst.zero_rat\, \<^const_name>\Nitpick.zero_frac\), (\<^const_name>\one_rat_inst.one_rat\, \<^const_name>\Nitpick.one_frac\), (\<^const_name>\plus_rat_inst.plus_rat\, \<^const_name>\Nitpick.plus_frac\), (\<^const_name>\times_rat_inst.times_rat\, \<^const_name>\Nitpick.times_frac\), (\<^const_name>\uminus_rat_inst.uminus_rat\, \<^const_name>\Nitpick.uminus_frac\), (\<^const_name>\inverse_rat_inst.inverse_rat\, \<^const_name>\Nitpick.inverse_frac\), (\<^const_name>\ord_rat_inst.less_rat\, \<^const_name>\Nitpick.less_frac\), (\<^const_name>\ord_rat_inst.less_eq_rat\, \<^const_name>\Nitpick.less_eq_frac\), (\<^const_name>\field_char_0_class.of_rat\, \<^const_name>\Nitpick.of_frac\)] \ lemmas [nitpick_unfold] = inverse_rat_inst.inverse_rat one_rat_inst.one_rat ord_rat_inst.less_rat ord_rat_inst.less_eq_rat plus_rat_inst.plus_rat times_rat_inst.times_rat uminus_rat_inst.uminus_rat zero_rat_inst.zero_rat subsection \Float syntax\ syntax "_Float" :: "float_const \ 'a" ("_") parse_translation \ let fun mk_frac str = let val {mant = i, exp = n} = Lexicon.read_float str; val exp = Syntax.const \<^const_syntax>\Power.power\; val ten = Numeral.mk_number_syntax 10; val exp10 = if n = 1 then ten else exp $ ten $ Numeral.mk_number_syntax n; in Syntax.const \<^const_syntax>\Fields.inverse_divide\ $ Numeral.mk_number_syntax i $ exp10 end; fun float_tr [(c as Const (\<^syntax_const>\_constrain\, _)) $ t $ u] = c $ float_tr [t] $ u | float_tr [t as Const (str, _)] = mk_frac str | float_tr ts = raise TERM ("float_tr", ts); in [(\<^syntax_const>\_Float\, K float_tr)] end \ text\Test:\ lemma "123.456 = -111.111 + 200 + 30 + 4 + 5/10 + 6/100 + (7/1000::rat)" by simp subsection \Hiding implementation details\ hide_const (open) normalize positive lifting_update rat.lifting lifting_forget rat.lifting end diff --git a/src/HOL/Real_Vector_Spaces.thy b/src/HOL/Real_Vector_Spaces.thy --- a/src/HOL/Real_Vector_Spaces.thy +++ b/src/HOL/Real_Vector_Spaces.thy @@ -1,2382 +1,2399 @@ (* Title: HOL/Real_Vector_Spaces.thy Author: Brian Huffman Author: Johannes Hölzl *) section \Vector Spaces and Algebras over the Reals\ theory Real_Vector_Spaces imports Real Topological_Spaces Vector_Spaces begin subsection \Real vector spaces\ class scaleR = fixes scaleR :: "real \ 'a \ 'a" (infixr "*\<^sub>R" 75) begin abbreviation divideR :: "'a \ real \ 'a" (infixl "'/\<^sub>R" 70) where "x /\<^sub>R r \ inverse r *\<^sub>R x" end class real_vector = scaleR + ab_group_add + assumes scaleR_add_right: "a *\<^sub>R (x + y) = a *\<^sub>R x + a *\<^sub>R y" and scaleR_add_left: "(a + b) *\<^sub>R x = a *\<^sub>R x + b *\<^sub>R x" and scaleR_scaleR: "a *\<^sub>R b *\<^sub>R x = (a * b) *\<^sub>R x" and scaleR_one: "1 *\<^sub>R x = x" class real_algebra = real_vector + ring + assumes mult_scaleR_left [simp]: "a *\<^sub>R x * y = a *\<^sub>R (x * y)" and mult_scaleR_right [simp]: "x * a *\<^sub>R y = a *\<^sub>R (x * y)" class real_algebra_1 = real_algebra + ring_1 class real_div_algebra = real_algebra_1 + division_ring class real_field = real_div_algebra + field instantiation real :: real_field begin definition real_scaleR_def [simp]: "scaleR a x = a * x" instance by standard (simp_all add: algebra_simps) end locale linear = Vector_Spaces.linear "scaleR::_\_\'a::real_vector" "scaleR::_\_\'b::real_vector" begin lemmas scaleR = scale end global_interpretation real_vector?: vector_space "scaleR :: real \ 'a \ 'a :: real_vector" rewrites "Vector_Spaces.linear (*\<^sub>R) (*\<^sub>R) = linear" and "Vector_Spaces.linear (*) (*\<^sub>R) = linear" defines dependent_raw_def: dependent = real_vector.dependent and representation_raw_def: representation = real_vector.representation and subspace_raw_def: subspace = real_vector.subspace and span_raw_def: span = real_vector.span and extend_basis_raw_def: extend_basis = real_vector.extend_basis and dim_raw_def: dim = real_vector.dim proof unfold_locales show "Vector_Spaces.linear (*\<^sub>R) (*\<^sub>R) = linear" "Vector_Spaces.linear (*) (*\<^sub>R) = linear" by (force simp: linear_def real_scaleR_def[abs_def])+ qed (use scaleR_add_right scaleR_add_left scaleR_scaleR scaleR_one in auto) hide_const (open)\ \locale constants\ real_vector.dependent real_vector.independent real_vector.representation real_vector.subspace real_vector.span real_vector.extend_basis real_vector.dim abbreviation "independent x \ \ dependent x" global_interpretation real_vector?: vector_space_pair "scaleR::_\_\'a::real_vector" "scaleR::_\_\'b::real_vector" rewrites "Vector_Spaces.linear (*\<^sub>R) (*\<^sub>R) = linear" and "Vector_Spaces.linear (*) (*\<^sub>R) = linear" defines construct_raw_def: construct = real_vector.construct proof unfold_locales show "Vector_Spaces.linear (*) (*\<^sub>R) = linear" unfolding linear_def real_scaleR_def by auto qed (auto simp: linear_def) hide_const (open)\ \locale constants\ real_vector.construct lemma linear_compose: "linear f \ linear g \ linear (g \ f)" unfolding linear_def by (rule Vector_Spaces.linear_compose) text \Recover original theorem names\ lemmas scaleR_left_commute = real_vector.scale_left_commute lemmas scaleR_zero_left = real_vector.scale_zero_left lemmas scaleR_minus_left = real_vector.scale_minus_left lemmas scaleR_diff_left = real_vector.scale_left_diff_distrib lemmas scaleR_sum_left = real_vector.scale_sum_left lemmas scaleR_zero_right = real_vector.scale_zero_right lemmas scaleR_minus_right = real_vector.scale_minus_right lemmas scaleR_diff_right = real_vector.scale_right_diff_distrib lemmas scaleR_sum_right = real_vector.scale_sum_right lemmas scaleR_eq_0_iff = real_vector.scale_eq_0_iff lemmas scaleR_left_imp_eq = real_vector.scale_left_imp_eq lemmas scaleR_right_imp_eq = real_vector.scale_right_imp_eq lemmas scaleR_cancel_left = real_vector.scale_cancel_left lemmas scaleR_cancel_right = real_vector.scale_cancel_right lemma [field_simps]: "c \ 0 \ a = b /\<^sub>R c \ c *\<^sub>R a = b" "c \ 0 \ b /\<^sub>R c = a \ b = c *\<^sub>R a" "c \ 0 \ a + b /\<^sub>R c = (c *\<^sub>R a + b) /\<^sub>R c" "c \ 0 \ a /\<^sub>R c + b = (a + c *\<^sub>R b) /\<^sub>R c" "c \ 0 \ a - b /\<^sub>R c = (c *\<^sub>R a - b) /\<^sub>R c" "c \ 0 \ a /\<^sub>R c - b = (a - c *\<^sub>R b) /\<^sub>R c" "c \ 0 \ - (a /\<^sub>R c) + b = (- a + c *\<^sub>R b) /\<^sub>R c" "c \ 0 \ - (a /\<^sub>R c) - b = (- a - c *\<^sub>R b) /\<^sub>R c" for a b :: "'a :: real_vector" by (auto simp add: scaleR_add_right scaleR_add_left scaleR_diff_right scaleR_diff_left) text \Legacy names\ lemmas scaleR_left_distrib = scaleR_add_left lemmas scaleR_right_distrib = scaleR_add_right lemmas scaleR_left_diff_distrib = scaleR_diff_left lemmas scaleR_right_diff_distrib = scaleR_diff_right lemmas linear_injective_0 = linear_inj_iff_eq_0 and linear_injective_on_subspace_0 = linear_inj_on_iff_eq_0 and linear_cmul = linear_scale and linear_scaleR = linear_scale_self and subspace_mul = subspace_scale and span_linear_image = linear_span_image and span_0 = span_zero and span_mul = span_scale and injective_scaleR = injective_scale lemma scaleR_minus1_left [simp]: "scaleR (-1) x = - x" for x :: "'a::real_vector" using scaleR_minus_left [of 1 x] by simp lemma scaleR_2: fixes x :: "'a::real_vector" shows "scaleR 2 x = x + x" unfolding one_add_one [symmetric] scaleR_left_distrib by simp lemma scaleR_half_double [simp]: fixes a :: "'a::real_vector" shows "(1 / 2) *\<^sub>R (a + a) = a" proof - have "\r. r *\<^sub>R (a + a) = (r * 2) *\<^sub>R a" by (metis scaleR_2 scaleR_scaleR) then show ?thesis by simp qed lemma linear_scale_real: fixes r::real shows "linear f \ f (r * b) = r * f b" using linear_scale by fastforce interpretation scaleR_left: additive "(\a. scaleR a x :: 'a::real_vector)" by standard (rule scaleR_left_distrib) interpretation scaleR_right: additive "(\x. scaleR a x :: 'a::real_vector)" by standard (rule scaleR_right_distrib) lemma nonzero_inverse_scaleR_distrib: "a \ 0 \ x \ 0 \ inverse (scaleR a x) = scaleR (inverse a) (inverse x)" for x :: "'a::real_div_algebra" by (rule inverse_unique) simp lemma inverse_scaleR_distrib: "inverse (scaleR a x) = scaleR (inverse a) (inverse x)" for x :: "'a::{real_div_algebra,division_ring}" by (metis inverse_zero nonzero_inverse_scaleR_distrib scale_eq_0_iff) lemmas sum_constant_scaleR = real_vector.sum_constant_scale\ \legacy name\ named_theorems vector_add_divide_simps "to simplify sums of scaled vectors" lemma [vector_add_divide_simps]: "v + (b / z) *\<^sub>R w = (if z = 0 then v else (z *\<^sub>R v + b *\<^sub>R w) /\<^sub>R z)" "a *\<^sub>R v + (b / z) *\<^sub>R w = (if z = 0 then a *\<^sub>R v else ((a * z) *\<^sub>R v + b *\<^sub>R w) /\<^sub>R z)" "(a / z) *\<^sub>R v + w = (if z = 0 then w else (a *\<^sub>R v + z *\<^sub>R w) /\<^sub>R z)" "(a / z) *\<^sub>R v + b *\<^sub>R w = (if z = 0 then b *\<^sub>R w else (a *\<^sub>R v + (b * z) *\<^sub>R w) /\<^sub>R z)" "v - (b / z) *\<^sub>R w = (if z = 0 then v else (z *\<^sub>R v - b *\<^sub>R w) /\<^sub>R z)" "a *\<^sub>R v - (b / z) *\<^sub>R w = (if z = 0 then a *\<^sub>R v else ((a * z) *\<^sub>R v - b *\<^sub>R w) /\<^sub>R z)" "(a / z) *\<^sub>R v - w = (if z = 0 then -w else (a *\<^sub>R v - z *\<^sub>R w) /\<^sub>R z)" "(a / z) *\<^sub>R v - b *\<^sub>R w = (if z = 0 then -b *\<^sub>R w else (a *\<^sub>R v - (b * z) *\<^sub>R w) /\<^sub>R z)" for v :: "'a :: real_vector" by (simp_all add: divide_inverse_commute scaleR_add_right scaleR_diff_right) lemma eq_vector_fraction_iff [vector_add_divide_simps]: fixes x :: "'a :: real_vector" shows "(x = (u / v) *\<^sub>R a) \ (if v=0 then x = 0 else v *\<^sub>R x = u *\<^sub>R a)" by auto (metis (no_types) divide_eq_1_iff divide_inverse_commute scaleR_one scaleR_scaleR) lemma vector_fraction_eq_iff [vector_add_divide_simps]: fixes x :: "'a :: real_vector" shows "((u / v) *\<^sub>R a = x) \ (if v=0 then x = 0 else u *\<^sub>R a = v *\<^sub>R x)" by (metis eq_vector_fraction_iff) lemma real_vector_affinity_eq: fixes x :: "'a :: real_vector" assumes m0: "m \ 0" shows "m *\<^sub>R x + c = y \ x = inverse m *\<^sub>R y - (inverse m *\<^sub>R c)" (is "?lhs \ ?rhs") proof assume ?lhs then have "m *\<^sub>R x = y - c" by (simp add: field_simps) then have "inverse m *\<^sub>R (m *\<^sub>R x) = inverse m *\<^sub>R (y - c)" by simp then show "x = inverse m *\<^sub>R y - (inverse m *\<^sub>R c)" using m0 by (simp add: scaleR_diff_right) next assume ?rhs with m0 show "m *\<^sub>R x + c = y" by (simp add: scaleR_diff_right) qed lemma real_vector_eq_affinity: "m \ 0 \ y = m *\<^sub>R x + c \ inverse m *\<^sub>R y - (inverse m *\<^sub>R c) = x" for x :: "'a::real_vector" using real_vector_affinity_eq[where m=m and x=x and y=y and c=c] by metis lemma scaleR_eq_iff [simp]: "b + u *\<^sub>R a = a + u *\<^sub>R b \ a = b \ u = 1" for a :: "'a::real_vector" proof (cases "u = 1") case True then show ?thesis by auto next case False have "a = b" if "b + u *\<^sub>R a = a + u *\<^sub>R b" proof - from that have "(u - 1) *\<^sub>R a = (u - 1) *\<^sub>R b" by (simp add: algebra_simps) with False show ?thesis by auto qed then show ?thesis by auto qed lemma scaleR_collapse [simp]: "(1 - u) *\<^sub>R a + u *\<^sub>R a = a" for a :: "'a::real_vector" by (simp add: algebra_simps) subsection \Embedding of the Reals into any \real_algebra_1\: \of_real\\ definition of_real :: "real \ 'a::real_algebra_1" where "of_real r = scaleR r 1" lemma scaleR_conv_of_real: "scaleR r x = of_real r * x" by (simp add: of_real_def) lemma of_real_0 [simp]: "of_real 0 = 0" by (simp add: of_real_def) lemma of_real_1 [simp]: "of_real 1 = 1" by (simp add: of_real_def) lemma of_real_add [simp]: "of_real (x + y) = of_real x + of_real y" by (simp add: of_real_def scaleR_left_distrib) lemma of_real_minus [simp]: "of_real (- x) = - of_real x" by (simp add: of_real_def) lemma of_real_diff [simp]: "of_real (x - y) = of_real x - of_real y" by (simp add: of_real_def scaleR_left_diff_distrib) lemma of_real_mult [simp]: "of_real (x * y) = of_real x * of_real y" by (simp add: of_real_def) lemma of_real_sum[simp]: "of_real (sum f s) = (\x\s. of_real (f x))" by (induct s rule: infinite_finite_induct) auto lemma of_real_prod[simp]: "of_real (prod f s) = (\x\s. of_real (f x))" by (induct s rule: infinite_finite_induct) auto lemma nonzero_of_real_inverse: "x \ 0 \ of_real (inverse x) = inverse (of_real x :: 'a::real_div_algebra)" by (simp add: of_real_def nonzero_inverse_scaleR_distrib) lemma of_real_inverse [simp]: "of_real (inverse x) = inverse (of_real x :: 'a::{real_div_algebra,division_ring})" by (simp add: of_real_def inverse_scaleR_distrib) lemma nonzero_of_real_divide: "y \ 0 \ of_real (x / y) = (of_real x / of_real y :: 'a::real_field)" by (simp add: divide_inverse nonzero_of_real_inverse) lemma of_real_divide [simp]: "of_real (x / y) = (of_real x / of_real y :: 'a::real_div_algebra)" by (simp add: divide_inverse) lemma of_real_power [simp]: "of_real (x ^ n) = (of_real x :: 'a::{real_algebra_1}) ^ n" by (induct n) simp_all lemma of_real_power_int [simp]: "of_real (power_int x n) = power_int (of_real x :: 'a :: {real_div_algebra,division_ring}) n" by (auto simp: power_int_def) lemma of_real_eq_iff [simp]: "of_real x = of_real y \ x = y" by (simp add: of_real_def) lemma inj_of_real: "inj of_real" by (auto intro: injI) lemmas of_real_eq_0_iff [simp] = of_real_eq_iff [of _ 0, simplified] lemmas of_real_eq_1_iff [simp] = of_real_eq_iff [of _ 1, simplified] lemma minus_of_real_eq_of_real_iff [simp]: "-of_real x = of_real y \ -x = y" using of_real_eq_iff[of "-x" y] by (simp only: of_real_minus) lemma of_real_eq_minus_of_real_iff [simp]: "of_real x = -of_real y \ x = -y" using of_real_eq_iff[of x "-y"] by (simp only: of_real_minus) lemma of_real_eq_id [simp]: "of_real = (id :: real \ real)" by (rule ext) (simp add: of_real_def) text \Collapse nested embeddings.\ lemma of_real_of_nat_eq [simp]: "of_real (of_nat n) = of_nat n" by (induct n) auto lemma of_real_of_int_eq [simp]: "of_real (of_int z) = of_int z" by (cases z rule: int_diff_cases) simp lemma of_real_numeral [simp]: "of_real (numeral w) = numeral w" using of_real_of_int_eq [of "numeral w"] by simp lemma of_real_neg_numeral [simp]: "of_real (- numeral w) = - numeral w" using of_real_of_int_eq [of "- numeral w"] by simp lemma numeral_power_int_eq_of_real_cancel_iff [simp]: "power_int (numeral x) n = (of_real y :: 'a :: {real_div_algebra, division_ring}) \ power_int (numeral x) n = y" proof - have "power_int (numeral x) n = (of_real (power_int (numeral x) n) :: 'a)" by simp also have "\ = of_real y \ power_int (numeral x) n = y" by (subst of_real_eq_iff) auto finally show ?thesis . qed lemma of_real_eq_numeral_power_int_cancel_iff [simp]: "(of_real y :: 'a :: {real_div_algebra, division_ring}) = power_int (numeral x) n \ y = power_int (numeral x) n" by (subst (1 2) eq_commute) simp lemma of_real_eq_of_real_power_int_cancel_iff [simp]: "power_int (of_real b :: 'a :: {real_div_algebra, division_ring}) w = of_real x \ power_int b w = x" by (metis of_real_power_int of_real_eq_iff) +lemma of_real_in_Ints_iff [simp]: "of_real x \ \ \ x \ \" +proof safe + fix x assume "(of_real x :: 'a) \ \" + then obtain n where "(of_real x :: 'a) = of_int n" + by (auto simp: Ints_def) + also have "of_int n = of_real (real_of_int n)" + by simp + finally have "x = real_of_int n" + by (subst (asm) of_real_eq_iff) + thus "x \ \" + by auto +qed (auto simp: Ints_def) + +lemma Ints_of_real [intro]: "x \ \ \ of_real x \ \" + by simp + + text \Every real algebra has characteristic zero.\ instance real_algebra_1 < ring_char_0 proof from inj_of_real inj_of_nat have "inj (of_real \ of_nat)" by (rule inj_compose) then show "inj (of_nat :: nat \ 'a)" by (simp add: comp_def) qed lemma fraction_scaleR_times [simp]: fixes a :: "'a::real_algebra_1" shows "(numeral u / numeral v) *\<^sub>R (numeral w * a) = (numeral u * numeral w / numeral v) *\<^sub>R a" by (metis (no_types, lifting) of_real_numeral scaleR_conv_of_real scaleR_scaleR times_divide_eq_left) lemma inverse_scaleR_times [simp]: fixes a :: "'a::real_algebra_1" shows "(1 / numeral v) *\<^sub>R (numeral w * a) = (numeral w / numeral v) *\<^sub>R a" by (metis divide_inverse_commute inverse_eq_divide of_real_numeral scaleR_conv_of_real scaleR_scaleR) lemma scaleR_times [simp]: fixes a :: "'a::real_algebra_1" shows "(numeral u) *\<^sub>R (numeral w * a) = (numeral u * numeral w) *\<^sub>R a" by (simp add: scaleR_conv_of_real) instance real_field < field_char_0 .. subsection \The Set of Real Numbers\ definition Reals :: "'a::real_algebra_1 set" ("\") where "\ = range of_real" lemma Reals_of_real [simp]: "of_real r \ \" by (simp add: Reals_def) lemma Reals_of_int [simp]: "of_int z \ \" by (subst of_real_of_int_eq [symmetric], rule Reals_of_real) lemma Reals_of_nat [simp]: "of_nat n \ \" by (subst of_real_of_nat_eq [symmetric], rule Reals_of_real) lemma Reals_numeral [simp]: "numeral w \ \" by (subst of_real_numeral [symmetric], rule Reals_of_real) lemma Reals_0 [simp]: "0 \ \" and Reals_1 [simp]: "1 \ \" by (simp_all add: Reals_def) lemma Reals_add [simp]: "a \ \ \ b \ \ \ a + b \ \" by (metis (no_types, hide_lams) Reals_def Reals_of_real imageE of_real_add) lemma Reals_minus [simp]: "a \ \ \ - a \ \" by (auto simp: Reals_def) lemma Reals_minus_iff [simp]: "- a \ \ \ a \ \" using Reals_minus by fastforce lemma Reals_diff [simp]: "a \ \ \ b \ \ \ a - b \ \" by (metis Reals_add Reals_minus_iff add_uminus_conv_diff) lemma Reals_mult [simp]: "a \ \ \ b \ \ \ a * b \ \" by (metis (no_types, lifting) Reals_def Reals_of_real imageE of_real_mult) lemma nonzero_Reals_inverse: "a \ \ \ a \ 0 \ inverse a \ \" for a :: "'a::real_div_algebra" by (metis Reals_def Reals_of_real imageE of_real_inverse) lemma Reals_inverse: "a \ \ \ inverse a \ \" for a :: "'a::{real_div_algebra,division_ring}" using nonzero_Reals_inverse by fastforce lemma Reals_inverse_iff [simp]: "inverse x \ \ \ x \ \" for x :: "'a::{real_div_algebra,division_ring}" by (metis Reals_inverse inverse_inverse_eq) lemma nonzero_Reals_divide: "a \ \ \ b \ \ \ b \ 0 \ a / b \ \" for a b :: "'a::real_field" by (simp add: divide_inverse) lemma Reals_divide [simp]: "a \ \ \ b \ \ \ a / b \ \" for a b :: "'a::{real_field,field}" using nonzero_Reals_divide by fastforce lemma Reals_power [simp]: "a \ \ \ a ^ n \ \" for a :: "'a::real_algebra_1" by (metis Reals_def Reals_of_real imageE of_real_power) lemma Reals_cases [cases set: Reals]: assumes "q \ \" obtains (of_real) r where "q = of_real r" unfolding Reals_def proof - from \q \ \\ have "q \ range of_real" unfolding Reals_def . then obtain r where "q = of_real r" .. then show thesis .. qed lemma sum_in_Reals [intro,simp]: "(\i. i \ s \ f i \ \) \ sum f s \ \" proof (induct s rule: infinite_finite_induct) case infinite then show ?case by (metis Reals_0 sum.infinite) qed simp_all lemma prod_in_Reals [intro,simp]: "(\i. i \ s \ f i \ \) \ prod f s \ \" proof (induct s rule: infinite_finite_induct) case infinite then show ?case by (metis Reals_1 prod.infinite) qed simp_all lemma Reals_induct [case_names of_real, induct set: Reals]: "q \ \ \ (\r. P (of_real r)) \ P q" by (rule Reals_cases) auto subsection \Ordered real vector spaces\ class ordered_real_vector = real_vector + ordered_ab_group_add + assumes scaleR_left_mono: "x \ y \ 0 \ a \ a *\<^sub>R x \ a *\<^sub>R y" and scaleR_right_mono: "a \ b \ 0 \ x \ a *\<^sub>R x \ b *\<^sub>R x" begin lemma scaleR_mono: "a \ b \ x \ y \ 0 \ b \ 0 \ x \ a *\<^sub>R x \ b *\<^sub>R y" by (meson order_trans scaleR_left_mono scaleR_right_mono) lemma scaleR_mono': "a \ b \ c \ d \ 0 \ a \ 0 \ c \ a *\<^sub>R c \ b *\<^sub>R d" by (rule scaleR_mono) (auto intro: order.trans) lemma pos_le_divideR_eq [field_simps]: "a \ b /\<^sub>R c \ c *\<^sub>R a \ b" (is "?P \ ?Q") if "0 < c" proof assume ?P with scaleR_left_mono that have "c *\<^sub>R a \ c *\<^sub>R (b /\<^sub>R c)" by simp with that show ?Q by (simp add: scaleR_one scaleR_scaleR inverse_eq_divide) next assume ?Q with scaleR_left_mono that have "c *\<^sub>R a /\<^sub>R c \ b /\<^sub>R c" by simp with that show ?P by (simp add: scaleR_one scaleR_scaleR inverse_eq_divide) qed lemma pos_less_divideR_eq [field_simps]: "a < b /\<^sub>R c \ c *\<^sub>R a < b" if "c > 0" using that pos_le_divideR_eq [of c a b] by (auto simp add: le_less scaleR_scaleR scaleR_one) lemma pos_divideR_le_eq [field_simps]: "b /\<^sub>R c \ a \ b \ c *\<^sub>R a" if "c > 0" using that pos_le_divideR_eq [of "inverse c" b a] by simp lemma pos_divideR_less_eq [field_simps]: "b /\<^sub>R c < a \ b < c *\<^sub>R a" if "c > 0" using that pos_less_divideR_eq [of "inverse c" b a] by simp lemma pos_le_minus_divideR_eq [field_simps]: "a \ - (b /\<^sub>R c) \ c *\<^sub>R a \ - b" if "c > 0" using that by (metis add_minus_cancel diff_0 left_minus minus_minus neg_le_iff_le scaleR_add_right uminus_add_conv_diff pos_le_divideR_eq) lemma pos_less_minus_divideR_eq [field_simps]: "a < - (b /\<^sub>R c) \ c *\<^sub>R a < - b" if "c > 0" using that by (metis le_less less_le_not_le pos_divideR_le_eq pos_divideR_less_eq pos_le_minus_divideR_eq) lemma pos_minus_divideR_le_eq [field_simps]: "- (b /\<^sub>R c) \ a \ - b \ c *\<^sub>R a" if "c > 0" using that by (metis pos_divideR_le_eq pos_le_minus_divideR_eq that inverse_positive_iff_positive le_imp_neg_le minus_minus) lemma pos_minus_divideR_less_eq [field_simps]: "- (b /\<^sub>R c) < a \ - b < c *\<^sub>R a" if "c > 0" using that by (simp add: less_le_not_le pos_le_minus_divideR_eq pos_minus_divideR_le_eq) lemma scaleR_image_atLeastAtMost: "c > 0 \ scaleR c ` {x..y} = {c *\<^sub>R x..c *\<^sub>R y}" apply (auto intro!: scaleR_left_mono simp: image_iff Bex_def) by (meson local.eq_iff pos_divideR_le_eq pos_le_divideR_eq) end lemma neg_le_divideR_eq [field_simps]: "a \ b /\<^sub>R c \ b \ c *\<^sub>R a" (is "?P \ ?Q") if "c < 0" for a b :: "'a :: ordered_real_vector" using that pos_le_divideR_eq [of "- c" a "- b"] by simp lemma neg_less_divideR_eq [field_simps]: "a < b /\<^sub>R c \ b < c *\<^sub>R a" if "c < 0" for a b :: "'a :: ordered_real_vector" using that neg_le_divideR_eq [of c a b] by (auto simp add: le_less) lemma neg_divideR_le_eq [field_simps]: "b /\<^sub>R c \ a \ c *\<^sub>R a \ b" if "c < 0" for a b :: "'a :: ordered_real_vector" using that pos_divideR_le_eq [of "- c" "- b" a] by simp lemma neg_divideR_less_eq [field_simps]: "b /\<^sub>R c < a \ c *\<^sub>R a < b" if "c < 0" for a b :: "'a :: ordered_real_vector" using that neg_divideR_le_eq [of c b a] by (auto simp add: le_less) lemma neg_le_minus_divideR_eq [field_simps]: "a \ - (b /\<^sub>R c) \ - b \ c *\<^sub>R a" if "c < 0" for a b :: "'a :: ordered_real_vector" using that pos_le_minus_divideR_eq [of "- c" a "- b"] by (simp add: minus_le_iff) lemma neg_less_minus_divideR_eq [field_simps]: "a < - (b /\<^sub>R c) \ - b < c *\<^sub>R a" if "c < 0" for a b :: "'a :: ordered_real_vector" proof - have *: "- b = c *\<^sub>R a \ b = - (c *\<^sub>R a)" by (metis add.inverse_inverse) from that neg_le_minus_divideR_eq [of c a b] show ?thesis by (auto simp add: le_less *) qed lemma neg_minus_divideR_le_eq [field_simps]: "- (b /\<^sub>R c) \ a \ c *\<^sub>R a \ - b" if "c < 0" for a b :: "'a :: ordered_real_vector" using that pos_minus_divideR_le_eq [of "- c" "- b" a] by (simp add: le_minus_iff) lemma neg_minus_divideR_less_eq [field_simps]: "- (b /\<^sub>R c) < a \ c *\<^sub>R a < - b" if "c < 0" for a b :: "'a :: ordered_real_vector" using that by (simp add: less_le_not_le neg_le_minus_divideR_eq neg_minus_divideR_le_eq) lemma [field_split_simps]: "a = b /\<^sub>R c \ (if c = 0 then a = 0 else c *\<^sub>R a = b)" "b /\<^sub>R c = a \ (if c = 0 then a = 0 else b = c *\<^sub>R a)" "a + b /\<^sub>R c = (if c = 0 then a else (c *\<^sub>R a + b) /\<^sub>R c)" "a /\<^sub>R c + b = (if c = 0 then b else (a + c *\<^sub>R b) /\<^sub>R c)" "a - b /\<^sub>R c = (if c = 0 then a else (c *\<^sub>R a - b) /\<^sub>R c)" "a /\<^sub>R c - b = (if c = 0 then - b else (a - c *\<^sub>R b) /\<^sub>R c)" "- (a /\<^sub>R c) + b = (if c = 0 then b else (- a + c *\<^sub>R b) /\<^sub>R c)" "- (a /\<^sub>R c) - b = (if c = 0 then - b else (- a - c *\<^sub>R b) /\<^sub>R c)" for a b :: "'a :: real_vector" by (auto simp add: field_simps) lemma [field_split_simps]: "0 < c \ a \ b /\<^sub>R c \ (if c > 0 then c *\<^sub>R a \ b else if c < 0 then b \ c *\<^sub>R a else a \ 0)" "0 < c \ a < b /\<^sub>R c \ (if c > 0 then c *\<^sub>R a < b else if c < 0 then b < c *\<^sub>R a else a < 0)" "0 < c \ b /\<^sub>R c \ a \ (if c > 0 then b \ c *\<^sub>R a else if c < 0 then c *\<^sub>R a \ b else a \ 0)" "0 < c \ b /\<^sub>R c < a \ (if c > 0 then b < c *\<^sub>R a else if c < 0 then c *\<^sub>R a < b else a > 0)" "0 < c \ a \ - (b /\<^sub>R c) \ (if c > 0 then c *\<^sub>R a \ - b else if c < 0 then - b \ c *\<^sub>R a else a \ 0)" "0 < c \ a < - (b /\<^sub>R c) \ (if c > 0 then c *\<^sub>R a < - b else if c < 0 then - b < c *\<^sub>R a else a < 0)" "0 < c \ - (b /\<^sub>R c) \ a \ (if c > 0 then - b \ c *\<^sub>R a else if c < 0 then c *\<^sub>R a \ - b else a \ 0)" "0 < c \ - (b /\<^sub>R c) < a \ (if c > 0 then - b < c *\<^sub>R a else if c < 0 then c *\<^sub>R a < - b else a > 0)" for a b :: "'a :: ordered_real_vector" by (clarsimp intro!: field_simps)+ lemma scaleR_nonneg_nonneg: "0 \ a \ 0 \ x \ 0 \ a *\<^sub>R x" for x :: "'a::ordered_real_vector" using scaleR_left_mono [of 0 x a] by simp lemma scaleR_nonneg_nonpos: "0 \ a \ x \ 0 \ a *\<^sub>R x \ 0" for x :: "'a::ordered_real_vector" using scaleR_left_mono [of x 0 a] by simp lemma scaleR_nonpos_nonneg: "a \ 0 \ 0 \ x \ a *\<^sub>R x \ 0" for x :: "'a::ordered_real_vector" using scaleR_right_mono [of a 0 x] by simp lemma split_scaleR_neg_le: "(0 \ a \ x \ 0) \ (a \ 0 \ 0 \ x) \ a *\<^sub>R x \ 0" for x :: "'a::ordered_real_vector" by (auto simp: scaleR_nonneg_nonpos scaleR_nonpos_nonneg) lemma le_add_iff1: "a *\<^sub>R e + c \ b *\<^sub>R e + d \ (a - b) *\<^sub>R e + c \ d" for c d e :: "'a::ordered_real_vector" by (simp add: algebra_simps) lemma le_add_iff2: "a *\<^sub>R e + c \ b *\<^sub>R e + d \ c \ (b - a) *\<^sub>R e + d" for c d e :: "'a::ordered_real_vector" by (simp add: algebra_simps) lemma scaleR_left_mono_neg: "b \ a \ c \ 0 \ c *\<^sub>R a \ c *\<^sub>R b" for a b :: "'a::ordered_real_vector" by (drule scaleR_left_mono [of _ _ "- c"], simp_all) lemma scaleR_right_mono_neg: "b \ a \ c \ 0 \ a *\<^sub>R c \ b *\<^sub>R c" for c :: "'a::ordered_real_vector" by (drule scaleR_right_mono [of _ _ "- c"], simp_all) lemma scaleR_nonpos_nonpos: "a \ 0 \ b \ 0 \ 0 \ a *\<^sub>R b" for b :: "'a::ordered_real_vector" using scaleR_right_mono_neg [of a 0 b] by simp lemma split_scaleR_pos_le: "(0 \ a \ 0 \ b) \ (a \ 0 \ b \ 0) \ 0 \ a *\<^sub>R b" for b :: "'a::ordered_real_vector" by (auto simp: scaleR_nonneg_nonneg scaleR_nonpos_nonpos) lemma zero_le_scaleR_iff: fixes b :: "'a::ordered_real_vector" shows "0 \ a *\<^sub>R b \ 0 < a \ 0 \ b \ a < 0 \ b \ 0 \ a = 0" (is "?lhs = ?rhs") proof (cases "a = 0") case True then show ?thesis by simp next case False show ?thesis proof assume ?lhs from \a \ 0\ consider "a > 0" | "a < 0" by arith then show ?rhs proof cases case 1 with \?lhs\ have "inverse a *\<^sub>R 0 \ inverse a *\<^sub>R (a *\<^sub>R b)" by (intro scaleR_mono) auto with 1 show ?thesis by simp next case 2 with \?lhs\ have "- inverse a *\<^sub>R 0 \ - inverse a *\<^sub>R (a *\<^sub>R b)" by (intro scaleR_mono) auto with 2 show ?thesis by simp qed next assume ?rhs then show ?lhs by (auto simp: not_le \a \ 0\ intro!: split_scaleR_pos_le) qed qed lemma scaleR_le_0_iff: "a *\<^sub>R b \ 0 \ 0 < a \ b \ 0 \ a < 0 \ 0 \ b \ a = 0" for b::"'a::ordered_real_vector" by (insert zero_le_scaleR_iff [of "-a" b]) force lemma scaleR_le_cancel_left: "c *\<^sub>R a \ c *\<^sub>R b \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" for b :: "'a::ordered_real_vector" by (auto simp: neq_iff scaleR_left_mono scaleR_left_mono_neg dest: scaleR_left_mono[where a="inverse c"] scaleR_left_mono_neg[where c="inverse c"]) lemma scaleR_le_cancel_left_pos: "0 < c \ c *\<^sub>R a \ c *\<^sub>R b \ a \ b" for b :: "'a::ordered_real_vector" by (auto simp: scaleR_le_cancel_left) lemma scaleR_le_cancel_left_neg: "c < 0 \ c *\<^sub>R a \ c *\<^sub>R b \ b \ a" for b :: "'a::ordered_real_vector" by (auto simp: scaleR_le_cancel_left) lemma scaleR_left_le_one_le: "0 \ x \ a \ 1 \ a *\<^sub>R x \ x" for x :: "'a::ordered_real_vector" and a :: real using scaleR_right_mono[of a 1 x] by simp subsection \Real normed vector spaces\ class dist = fixes dist :: "'a \ 'a \ real" class norm = fixes norm :: "'a \ real" class sgn_div_norm = scaleR + norm + sgn + assumes sgn_div_norm: "sgn x = x /\<^sub>R norm x" class dist_norm = dist + norm + minus + assumes dist_norm: "dist x y = norm (x - y)" class uniformity_dist = dist + uniformity + assumes uniformity_dist: "uniformity = (INF e\{0 <..}. principal {(x, y). dist x y < e})" begin lemma eventually_uniformity_metric: "eventually P uniformity \ (\e>0. \x y. dist x y < e \ P (x, y))" unfolding uniformity_dist by (subst eventually_INF_base) (auto simp: eventually_principal subset_eq intro: bexI[of _ "min _ _"]) end class real_normed_vector = real_vector + sgn_div_norm + dist_norm + uniformity_dist + open_uniformity + assumes norm_eq_zero [simp]: "norm x = 0 \ x = 0" and norm_triangle_ineq: "norm (x + y) \ norm x + norm y" and norm_scaleR [simp]: "norm (scaleR a x) = \a\ * norm x" begin lemma norm_ge_zero [simp]: "0 \ norm x" proof - have "0 = norm (x + -1 *\<^sub>R x)" using scaleR_add_left[of 1 "-1" x] norm_scaleR[of 0 x] by (simp add: scaleR_one) also have "\ \ norm x + norm (-1 *\<^sub>R x)" by (rule norm_triangle_ineq) finally show ?thesis by simp qed end class real_normed_algebra = real_algebra + real_normed_vector + assumes norm_mult_ineq: "norm (x * y) \ norm x * norm y" class real_normed_algebra_1 = real_algebra_1 + real_normed_algebra + assumes norm_one [simp]: "norm 1 = 1" lemma (in real_normed_algebra_1) scaleR_power [simp]: "(scaleR x y) ^ n = scaleR (x^n) (y^n)" by (induct n) (simp_all add: scaleR_one scaleR_scaleR mult_ac) class real_normed_div_algebra = real_div_algebra + real_normed_vector + assumes norm_mult: "norm (x * y) = norm x * norm y" class real_normed_field = real_field + real_normed_div_algebra instance real_normed_div_algebra < real_normed_algebra_1 proof show "norm (x * y) \ norm x * norm y" for x y :: 'a by (simp add: norm_mult) next have "norm (1 * 1::'a) = norm (1::'a) * norm (1::'a)" by (rule norm_mult) then show "norm (1::'a) = 1" by simp qed context real_normed_vector begin lemma norm_zero [simp]: "norm (0::'a) = 0" by simp lemma zero_less_norm_iff [simp]: "norm x > 0 \ x \ 0" by (simp add: order_less_le) lemma norm_not_less_zero [simp]: "\ norm x < 0" by (simp add: linorder_not_less) lemma norm_le_zero_iff [simp]: "norm x \ 0 \ x = 0" by (simp add: order_le_less) lemma norm_minus_cancel [simp]: "norm (- x) = norm x" proof - have "- 1 *\<^sub>R x = - (1 *\<^sub>R x)" unfolding add_eq_0_iff2[symmetric] scaleR_add_left[symmetric] using norm_eq_zero by fastforce then have "norm (- x) = norm (scaleR (- 1) x)" by (simp only: scaleR_one) also have "\ = \- 1\ * norm x" by (rule norm_scaleR) finally show ?thesis by simp qed lemma norm_minus_commute: "norm (a - b) = norm (b - a)" proof - have "norm (- (b - a)) = norm (b - a)" by (rule norm_minus_cancel) then show ?thesis by simp qed lemma dist_add_cancel [simp]: "dist (a + b) (a + c) = dist b c" by (simp add: dist_norm) lemma dist_add_cancel2 [simp]: "dist (b + a) (c + a) = dist b c" by (simp add: dist_norm) lemma norm_uminus_minus: "norm (- x - y) = norm (x + y)" by (subst (2) norm_minus_cancel[symmetric], subst minus_add_distrib) simp lemma norm_triangle_ineq2: "norm a - norm b \ norm (a - b)" proof - have "norm (a - b + b) \ norm (a - b) + norm b" by (rule norm_triangle_ineq) then show ?thesis by simp qed lemma norm_triangle_ineq3: "\norm a - norm b\ \ norm (a - b)" proof - have "norm a - norm b \ norm (a - b)" by (simp add: norm_triangle_ineq2) moreover have "norm b - norm a \ norm (a - b)" by (metis norm_minus_commute norm_triangle_ineq2) ultimately show ?thesis by (simp add: abs_le_iff) qed lemma norm_triangle_ineq4: "norm (a - b) \ norm a + norm b" proof - have "norm (a + - b) \ norm a + norm (- b)" by (rule norm_triangle_ineq) then show ?thesis by simp qed lemma norm_triangle_le_diff: "norm x + norm y \ e \ norm (x - y) \ e" by (meson norm_triangle_ineq4 order_trans) lemma norm_diff_ineq: "norm a - norm b \ norm (a + b)" proof - have "norm a - norm (- b) \ norm (a - - b)" by (rule norm_triangle_ineq2) then show ?thesis by simp qed lemma norm_triangle_sub: "norm x \ norm y + norm (x - y)" using norm_triangle_ineq[of "y" "x - y"] by (simp add: field_simps) lemma norm_triangle_le: "norm x + norm y \ e \ norm (x + y) \ e" by (rule norm_triangle_ineq [THEN order_trans]) lemma norm_triangle_lt: "norm x + norm y < e \ norm (x + y) < e" by (rule norm_triangle_ineq [THEN le_less_trans]) lemma norm_add_leD: "norm (a + b) \ c \ norm b \ norm a + c" by (metis ab_semigroup_add_class.add.commute add_commute diff_le_eq norm_diff_ineq order_trans) lemma norm_diff_triangle_ineq: "norm ((a + b) - (c + d)) \ norm (a - c) + norm (b - d)" proof - have "norm ((a + b) - (c + d)) = norm ((a - c) + (b - d))" by (simp add: algebra_simps) also have "\ \ norm (a - c) + norm (b - d)" by (rule norm_triangle_ineq) finally show ?thesis . qed lemma norm_diff_triangle_le: "norm (x - z) \ e1 + e2" if "norm (x - y) \ e1" "norm (y - z) \ e2" proof - have "norm (x - (y + z - y)) \ norm (x - y) + norm (y - z)" using norm_diff_triangle_ineq that diff_diff_eq2 by presburger with that show ?thesis by simp qed lemma norm_diff_triangle_less: "norm (x - z) < e1 + e2" if "norm (x - y) < e1" "norm (y - z) < e2" proof - have "norm (x - z) \ norm (x - y) + norm (y - z)" by (metis norm_diff_triangle_ineq add_diff_cancel_left' diff_diff_eq2) with that show ?thesis by auto qed lemma norm_triangle_mono: "norm a \ r \ norm b \ s \ norm (a + b) \ r + s" by (metis (mono_tags) add_mono_thms_linordered_semiring(1) norm_triangle_ineq order.trans) lemma norm_sum: "norm (sum f A) \ (\i\A. norm (f i))" for f::"'b \ 'a" by (induct A rule: infinite_finite_induct) (auto intro: norm_triangle_mono) lemma sum_norm_le: "norm (sum f S) \ sum g S" if "\x. x \ S \ norm (f x) \ g x" for f::"'b \ 'a" by (rule order_trans [OF norm_sum sum_mono]) (simp add: that) lemma abs_norm_cancel [simp]: "\norm a\ = norm a" by (rule abs_of_nonneg [OF norm_ge_zero]) lemma sum_norm_bound: "norm (sum f S) \ of_nat (card S)*K" if "\x. x \ S \ norm (f x) \ K" for f :: "'b \ 'a" using sum_norm_le[OF that] sum_constant[symmetric] by simp lemma norm_add_less: "norm x < r \ norm y < s \ norm (x + y) < r + s" by (rule order_le_less_trans [OF norm_triangle_ineq add_strict_mono]) end lemma dist_scaleR [simp]: "dist (x *\<^sub>R a) (y *\<^sub>R a) = \x - y\ * norm a" for a :: "'a::real_normed_vector" by (metis dist_norm norm_scaleR scaleR_left.diff) lemma norm_mult_less: "norm x < r \ norm y < s \ norm (x * y) < r * s" for x y :: "'a::real_normed_algebra" by (rule order_le_less_trans [OF norm_mult_ineq]) (simp add: mult_strict_mono') lemma norm_of_real [simp]: "norm (of_real r :: 'a::real_normed_algebra_1) = \r\" by (simp add: of_real_def) lemma norm_numeral [simp]: "norm (numeral w::'a::real_normed_algebra_1) = numeral w" by (subst of_real_numeral [symmetric], subst norm_of_real, simp) lemma norm_neg_numeral [simp]: "norm (- numeral w::'a::real_normed_algebra_1) = numeral w" by (subst of_real_neg_numeral [symmetric], subst norm_of_real, simp) lemma norm_of_real_add1 [simp]: "norm (of_real x + 1 :: 'a :: real_normed_div_algebra) = \x + 1\" by (metis norm_of_real of_real_1 of_real_add) lemma norm_of_real_addn [simp]: "norm (of_real x + numeral b :: 'a :: real_normed_div_algebra) = \x + numeral b\" by (metis norm_of_real of_real_add of_real_numeral) lemma norm_of_int [simp]: "norm (of_int z::'a::real_normed_algebra_1) = \of_int z\" by (subst of_real_of_int_eq [symmetric], rule norm_of_real) lemma norm_of_nat [simp]: "norm (of_nat n::'a::real_normed_algebra_1) = of_nat n" by (metis abs_of_nat norm_of_real of_real_of_nat_eq) lemma nonzero_norm_inverse: "a \ 0 \ norm (inverse a) = inverse (norm a)" for a :: "'a::real_normed_div_algebra" by (metis inverse_unique norm_mult norm_one right_inverse) lemma norm_inverse: "norm (inverse a) = inverse (norm a)" for a :: "'a::{real_normed_div_algebra,division_ring}" by (metis inverse_zero nonzero_norm_inverse norm_zero) lemma nonzero_norm_divide: "b \ 0 \ norm (a / b) = norm a / norm b" for a b :: "'a::real_normed_field" by (simp add: divide_inverse norm_mult nonzero_norm_inverse) lemma norm_divide: "norm (a / b) = norm a / norm b" for a b :: "'a::{real_normed_field,field}" by (simp add: divide_inverse norm_mult norm_inverse) lemma norm_inverse_le_norm: fixes x :: "'a::real_normed_div_algebra" shows "r \ norm x \ 0 < r \ norm (inverse x) \ inverse r" by (simp add: le_imp_inverse_le norm_inverse) lemma norm_power_ineq: "norm (x ^ n) \ norm x ^ n" for x :: "'a::real_normed_algebra_1" proof (induct n) case 0 show "norm (x ^ 0) \ norm x ^ 0" by simp next case (Suc n) have "norm (x * x ^ n) \ norm x * norm (x ^ n)" by (rule norm_mult_ineq) also from Suc have "\ \ norm x * norm x ^ n" using norm_ge_zero by (rule mult_left_mono) finally show "norm (x ^ Suc n) \ norm x ^ Suc n" by simp qed lemma norm_power: "norm (x ^ n) = norm x ^ n" for x :: "'a::real_normed_div_algebra" by (induct n) (simp_all add: norm_mult) lemma norm_power_int: "norm (power_int x n) = power_int (norm x) n" for x :: "'a::real_normed_div_algebra" by (cases n rule: int_cases4) (auto simp: norm_power power_int_minus norm_inverse) lemma power_eq_imp_eq_norm: fixes w :: "'a::real_normed_div_algebra" assumes eq: "w ^ n = z ^ n" and "n > 0" shows "norm w = norm z" proof - have "norm w ^ n = norm z ^ n" by (metis (no_types) eq norm_power) then show ?thesis using assms by (force intro: power_eq_imp_eq_base) qed lemma power_eq_1_iff: fixes w :: "'a::real_normed_div_algebra" shows "w ^ n = 1 \ norm w = 1 \ n = 0" by (metis norm_one power_0_left power_eq_0_iff power_eq_imp_eq_norm power_one) lemma norm_mult_numeral1 [simp]: "norm (numeral w * a) = numeral w * norm a" for a b :: "'a::{real_normed_field,field}" by (simp add: norm_mult) lemma norm_mult_numeral2 [simp]: "norm (a * numeral w) = norm a * numeral w" for a b :: "'a::{real_normed_field,field}" by (simp add: norm_mult) lemma norm_divide_numeral [simp]: "norm (a / numeral w) = norm a / numeral w" for a b :: "'a::{real_normed_field,field}" by (simp add: norm_divide) lemma norm_of_real_diff [simp]: "norm (of_real b - of_real a :: 'a::real_normed_algebra_1) \ \b - a\" by (metis norm_of_real of_real_diff order_refl) text \Despite a superficial resemblance, \norm_eq_1\ is not relevant.\ lemma square_norm_one: fixes x :: "'a::real_normed_div_algebra" assumes "x\<^sup>2 = 1" shows "norm x = 1" by (metis assms norm_minus_cancel norm_one power2_eq_1_iff) lemma norm_less_p1: "norm x < norm (of_real (norm x) + 1 :: 'a)" for x :: "'a::real_normed_algebra_1" proof - have "norm x < norm (of_real (norm x + 1) :: 'a)" by (simp add: of_real_def) then show ?thesis by simp qed lemma prod_norm: "prod (\x. norm (f x)) A = norm (prod f A)" for f :: "'a \ 'b::{comm_semiring_1,real_normed_div_algebra}" by (induct A rule: infinite_finite_induct) (auto simp: norm_mult) lemma norm_prod_le: "norm (prod f A) \ (\a\A. norm (f a :: 'a :: {real_normed_algebra_1,comm_monoid_mult}))" proof (induct A rule: infinite_finite_induct) case empty then show ?case by simp next case (insert a A) then have "norm (prod f (insert a A)) \ norm (f a) * norm (prod f A)" by (simp add: norm_mult_ineq) also have "norm (prod f A) \ (\a\A. norm (f a))" by (rule insert) finally show ?case by (simp add: insert mult_left_mono) next case infinite then show ?case by simp qed lemma norm_prod_diff: fixes z w :: "'i \ 'a::{real_normed_algebra_1, comm_monoid_mult}" shows "(\i. i \ I \ norm (z i) \ 1) \ (\i. i \ I \ norm (w i) \ 1) \ norm ((\i\I. z i) - (\i\I. w i)) \ (\i\I. norm (z i - w i))" proof (induction I rule: infinite_finite_induct) case empty then show ?case by simp next case (insert i I) note insert.hyps[simp] have "norm ((\i\insert i I. z i) - (\i\insert i I. w i)) = norm ((\i\I. z i) * (z i - w i) + ((\i\I. z i) - (\i\I. w i)) * w i)" (is "_ = norm (?t1 + ?t2)") by (auto simp: field_simps) also have "\ \ norm ?t1 + norm ?t2" by (rule norm_triangle_ineq) also have "norm ?t1 \ norm (\i\I. z i) * norm (z i - w i)" by (rule norm_mult_ineq) also have "\ \ (\i\I. norm (z i)) * norm(z i - w i)" by (rule mult_right_mono) (auto intro: norm_prod_le) also have "(\i\I. norm (z i)) \ (\i\I. 1)" by (intro prod_mono) (auto intro!: insert) also have "norm ?t2 \ norm ((\i\I. z i) - (\i\I. w i)) * norm (w i)" by (rule norm_mult_ineq) also have "norm (w i) \ 1" by (auto intro: insert) also have "norm ((\i\I. z i) - (\i\I. w i)) \ (\i\I. norm (z i - w i))" using insert by auto finally show ?case by (auto simp: ac_simps mult_right_mono mult_left_mono) next case infinite then show ?case by simp qed lemma norm_power_diff: fixes z w :: "'a::{real_normed_algebra_1, comm_monoid_mult}" assumes "norm z \ 1" "norm w \ 1" shows "norm (z^m - w^m) \ m * norm (z - w)" proof - have "norm (z^m - w^m) = norm ((\ i < m. z) - (\ i < m. w))" by simp also have "\ \ (\i = m * norm (z - w)" by simp finally show ?thesis . qed subsection \Metric spaces\ class metric_space = uniformity_dist + open_uniformity + assumes dist_eq_0_iff [simp]: "dist x y = 0 \ x = y" and dist_triangle2: "dist x y \ dist x z + dist y z" begin lemma dist_self [simp]: "dist x x = 0" by simp lemma zero_le_dist [simp]: "0 \ dist x y" using dist_triangle2 [of x x y] by simp lemma zero_less_dist_iff: "0 < dist x y \ x \ y" by (simp add: less_le) lemma dist_not_less_zero [simp]: "\ dist x y < 0" by (simp add: not_less) lemma dist_le_zero_iff [simp]: "dist x y \ 0 \ x = y" by (simp add: le_less) lemma dist_commute: "dist x y = dist y x" proof (rule order_antisym) show "dist x y \ dist y x" using dist_triangle2 [of x y x] by simp show "dist y x \ dist x y" using dist_triangle2 [of y x y] by simp qed lemma dist_commute_lessI: "dist y x < e \ dist x y < e" by (simp add: dist_commute) lemma dist_triangle: "dist x z \ dist x y + dist y z" using dist_triangle2 [of x z y] by (simp add: dist_commute) lemma dist_triangle3: "dist x y \ dist a x + dist a y" using dist_triangle2 [of x y a] by (simp add: dist_commute) lemma abs_dist_diff_le: "\dist a b - dist b c\ \ dist a c" using dist_triangle3[of b c a] dist_triangle2[of a b c] by simp lemma dist_pos_lt: "x \ y \ 0 < dist x y" by (simp add: zero_less_dist_iff) lemma dist_nz: "x \ y \ 0 < dist x y" by (simp add: zero_less_dist_iff) declare dist_nz [symmetric, simp] lemma dist_triangle_le: "dist x z + dist y z \ e \ dist x y \ e" by (rule order_trans [OF dist_triangle2]) lemma dist_triangle_lt: "dist x z + dist y z < e \ dist x y < e" by (rule le_less_trans [OF dist_triangle2]) lemma dist_triangle_less_add: "dist x1 y < e1 \ dist x2 y < e2 \ dist x1 x2 < e1 + e2" by (rule dist_triangle_lt [where z=y]) simp lemma dist_triangle_half_l: "dist x1 y < e / 2 \ dist x2 y < e / 2 \ dist x1 x2 < e" by (rule dist_triangle_lt [where z=y]) simp lemma dist_triangle_half_r: "dist y x1 < e / 2 \ dist y x2 < e / 2 \ dist x1 x2 < e" by (rule dist_triangle_half_l) (simp_all add: dist_commute) lemma dist_triangle_third: assumes "dist x1 x2 < e/3" "dist x2 x3 < e/3" "dist x3 x4 < e/3" shows "dist x1 x4 < e" proof - have "dist x1 x3 < e/3 + e/3" by (metis assms(1) assms(2) dist_commute dist_triangle_less_add) then have "dist x1 x4 < (e/3 + e/3) + e/3" by (metis assms(3) dist_commute dist_triangle_less_add) then show ?thesis by simp qed subclass uniform_space proof fix E x assume "eventually E uniformity" then obtain e where E: "0 < e" "\x y. dist x y < e \ E (x, y)" by (auto simp: eventually_uniformity_metric) then show "E (x, x)" "\\<^sub>F (x, y) in uniformity. E (y, x)" by (auto simp: eventually_uniformity_metric dist_commute) show "\D. eventually D uniformity \ (\x y z. D (x, y) \ D (y, z) \ E (x, z))" using E dist_triangle_half_l[where e=e] unfolding eventually_uniformity_metric by (intro exI[of _ "\(x, y). dist x y < e / 2"] exI[of _ "e/2"] conjI) (auto simp: dist_commute) qed lemma open_dist: "open S \ (\x\S. \e>0. \y. dist y x < e \ y \ S)" by (simp add: dist_commute open_uniformity eventually_uniformity_metric) lemma open_ball: "open {y. dist x y < d}" unfolding open_dist proof (intro ballI) fix y assume *: "y \ {y. dist x y < d}" then show "\e>0. \z. dist z y < e \ z \ {y. dist x y < d}" by (auto intro!: exI[of _ "d - dist x y"] simp: field_simps dist_triangle_lt) qed subclass first_countable_topology proof fix x show "\A::nat \ 'a set. (\i. x \ A i \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" proof (safe intro!: exI[of _ "\n. {y. dist x y < inverse (Suc n)}"]) fix S assume "open S" "x \ S" then obtain e where e: "0 < e" and "{y. dist x y < e} \ S" by (auto simp: open_dist subset_eq dist_commute) moreover from e obtain i where "inverse (Suc i) < e" by (auto dest!: reals_Archimedean) then have "{y. dist x y < inverse (Suc i)} \ {y. dist x y < e}" by auto ultimately show "\i. {y. dist x y < inverse (Suc i)} \ S" by blast qed (auto intro: open_ball) qed end instance metric_space \ t2_space proof fix x y :: "'a::metric_space" assume xy: "x \ y" let ?U = "{y'. dist x y' < dist x y / 2}" let ?V = "{x'. dist y x' < dist x y / 2}" have *: "d x z \ d x y + d y z \ d y z = d z y \ \ (d x y * 2 < d x z \ d z y * 2 < d x z)" for d :: "'a \ 'a \ real" and x y z :: 'a by arith have "open ?U \ open ?V \ x \ ?U \ y \ ?V \ ?U \ ?V = {}" using dist_pos_lt[OF xy] *[of dist, OF dist_triangle dist_commute] using open_ball[of _ "dist x y / 2"] by auto then show "\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" by blast qed text \Every normed vector space is a metric space.\ instance real_normed_vector < metric_space proof fix x y z :: 'a show "dist x y = 0 \ x = y" by (simp add: dist_norm) show "dist x y \ dist x z + dist y z" using norm_triangle_ineq4 [of "x - z" "y - z"] by (simp add: dist_norm) qed subsection \Class instances for real numbers\ instantiation real :: real_normed_field begin definition dist_real_def: "dist x y = \x - y\" definition uniformity_real_def [code del]: "(uniformity :: (real \ real) filter) = (INF e\{0 <..}. principal {(x, y). dist x y < e})" definition open_real_def [code del]: "open (U :: real set) \ (\x\U. eventually (\(x', y). x' = x \ y \ U) uniformity)" definition real_norm_def [simp]: "norm r = \r\" instance by intro_classes (auto simp: abs_mult open_real_def dist_real_def sgn_real_def uniformity_real_def) end declare uniformity_Abort[where 'a=real, code] lemma dist_of_real [simp]: "dist (of_real x :: 'a) (of_real y) = dist x y" for a :: "'a::real_normed_div_algebra" by (metis dist_norm norm_of_real of_real_diff real_norm_def) declare [[code abort: "open :: real set \ bool"]] instance real :: linorder_topology proof show "(open :: real set \ bool) = generate_topology (range lessThan \ range greaterThan)" proof (rule ext, safe) fix S :: "real set" assume "open S" then obtain f where "\x\S. 0 < f x \ (\y. dist y x < f x \ y \ S)" unfolding open_dist bchoice_iff .. then have *: "(\x\S. {x - f x <..} \ {..< x + f x}) = S" (is "?S = S") by (fastforce simp: dist_real_def) moreover have "generate_topology (range lessThan \ range greaterThan) ?S" by (force intro: generate_topology.Basis generate_topology_Union generate_topology.Int) ultimately show "generate_topology (range lessThan \ range greaterThan) S" by simp next fix S :: "real set" assume "generate_topology (range lessThan \ range greaterThan) S" moreover have "\a::real. open {.. (\y. \y - x\ < a - x \ y \ {..e>0. \y. \y - x\ < e \ y \ {..a::real. open {a <..}" unfolding open_dist dist_real_def proof clarify fix x a :: real assume "a < x" then have "0 < x - a \ (\y. \y - x\ < x - a \ y \ {a<..})" by auto then show "\e>0. \y. \y - x\ < e \ y \ {a<..}" .. qed ultimately show "open S" by induct auto qed qed instance real :: linear_continuum_topology .. lemmas open_real_greaterThan = open_greaterThan[where 'a=real] lemmas open_real_lessThan = open_lessThan[where 'a=real] lemmas open_real_greaterThanLessThan = open_greaterThanLessThan[where 'a=real] lemmas closed_real_atMost = closed_atMost[where 'a=real] lemmas closed_real_atLeast = closed_atLeast[where 'a=real] lemmas closed_real_atLeastAtMost = closed_atLeastAtMost[where 'a=real] instance real :: ordered_real_vector by standard (auto intro: mult_left_mono mult_right_mono) subsection \Extra type constraints\ text \Only allow \<^term>\open\ in class \topological_space\.\ setup \Sign.add_const_constraint (\<^const_name>\open\, SOME \<^typ>\'a::topological_space set \ bool\)\ text \Only allow \<^term>\uniformity\ in class \uniform_space\.\ setup \Sign.add_const_constraint (\<^const_name>\uniformity\, SOME \<^typ>\('a::uniformity \ 'a) filter\)\ text \Only allow \<^term>\dist\ in class \metric_space\.\ setup \Sign.add_const_constraint (\<^const_name>\dist\, SOME \<^typ>\'a::metric_space \ 'a \ real\)\ text \Only allow \<^term>\norm\ in class \real_normed_vector\.\ setup \Sign.add_const_constraint (\<^const_name>\norm\, SOME \<^typ>\'a::real_normed_vector \ real\)\ subsection \Sign function\ lemma norm_sgn: "norm (sgn x) = (if x = 0 then 0 else 1)" for x :: "'a::real_normed_vector" by (simp add: sgn_div_norm) lemma sgn_zero [simp]: "sgn (0::'a::real_normed_vector) = 0" by (simp add: sgn_div_norm) lemma sgn_zero_iff: "sgn x = 0 \ x = 0" for x :: "'a::real_normed_vector" by (simp add: sgn_div_norm) lemma sgn_minus: "sgn (- x) = - sgn x" for x :: "'a::real_normed_vector" by (simp add: sgn_div_norm) lemma sgn_scaleR: "sgn (scaleR r x) = scaleR (sgn r) (sgn x)" for x :: "'a::real_normed_vector" by (simp add: sgn_div_norm ac_simps) lemma sgn_one [simp]: "sgn (1::'a::real_normed_algebra_1) = 1" by (simp add: sgn_div_norm) lemma sgn_of_real: "sgn (of_real r :: 'a::real_normed_algebra_1) = of_real (sgn r)" unfolding of_real_def by (simp only: sgn_scaleR sgn_one) lemma sgn_mult: "sgn (x * y) = sgn x * sgn y" for x y :: "'a::real_normed_div_algebra" by (simp add: sgn_div_norm norm_mult) hide_fact (open) sgn_mult lemma real_sgn_eq: "sgn x = x / \x\" for x :: real by (simp add: sgn_div_norm divide_inverse) lemma zero_le_sgn_iff [simp]: "0 \ sgn x \ 0 \ x" for x :: real by (cases "0::real" x rule: linorder_cases) simp_all lemma sgn_le_0_iff [simp]: "sgn x \ 0 \ x \ 0" for x :: real by (cases "0::real" x rule: linorder_cases) simp_all lemma norm_conv_dist: "norm x = dist x 0" unfolding dist_norm by simp declare norm_conv_dist [symmetric, simp] lemma dist_0_norm [simp]: "dist 0 x = norm x" for x :: "'a::real_normed_vector" by (simp add: dist_norm) lemma dist_diff [simp]: "dist a (a - b) = norm b" "dist (a - b) a = norm b" by (simp_all add: dist_norm) lemma dist_of_int: "dist (of_int m) (of_int n :: 'a :: real_normed_algebra_1) = of_int \m - n\" proof - have "dist (of_int m) (of_int n :: 'a) = dist (of_int m :: 'a) (of_int m - (of_int (m - n)))" by simp also have "\ = of_int \m - n\" by (subst dist_diff, subst norm_of_int) simp finally show ?thesis . qed lemma dist_of_nat: "dist (of_nat m) (of_nat n :: 'a :: real_normed_algebra_1) = of_int \int m - int n\" by (subst (1 2) of_int_of_nat_eq [symmetric]) (rule dist_of_int) subsection \Bounded Linear and Bilinear Operators\ lemma linearI: "linear f" if "\b1 b2. f (b1 + b2) = f b1 + f b2" "\r b. f (r *\<^sub>R b) = r *\<^sub>R f b" using that by unfold_locales (auto simp: algebra_simps) lemma linear_iff: "linear f \ (\x y. f (x + y) = f x + f y) \ (\c x. f (c *\<^sub>R x) = c *\<^sub>R f x)" (is "linear f \ ?rhs") proof assume "linear f" then interpret f: linear f . show "?rhs" by (simp add: f.add f.scale) next assume "?rhs" then show "linear f" by (intro linearI) auto qed lemmas linear_scaleR_left = linear_scale_left lemmas linear_imp_scaleR = linear_imp_scale corollary real_linearD: fixes f :: "real \ real" assumes "linear f" obtains c where "f = (*) c" by (rule linear_imp_scaleR [OF assms]) (force simp: scaleR_conv_of_real) lemma linear_times_of_real: "linear (\x. a * of_real x)" by (auto intro!: linearI simp: distrib_left) (metis mult_scaleR_right scaleR_conv_of_real) locale bounded_linear = linear f for f :: "'a::real_normed_vector \ 'b::real_normed_vector" + assumes bounded: "\K. \x. norm (f x) \ norm x * K" begin lemma pos_bounded: "\K>0. \x. norm (f x) \ norm x * K" proof - obtain K where K: "\x. norm (f x) \ norm x * K" using bounded by blast show ?thesis proof (intro exI impI conjI allI) show "0 < max 1 K" by (rule order_less_le_trans [OF zero_less_one max.cobounded1]) next fix x have "norm (f x) \ norm x * K" using K . also have "\ \ norm x * max 1 K" by (rule mult_left_mono [OF max.cobounded2 norm_ge_zero]) finally show "norm (f x) \ norm x * max 1 K" . qed qed lemma nonneg_bounded: "\K\0. \x. norm (f x) \ norm x * K" using pos_bounded by (auto intro: order_less_imp_le) lemma linear: "linear f" by (fact local.linear_axioms) end lemma bounded_linear_intro: assumes "\x y. f (x + y) = f x + f y" and "\r x. f (scaleR r x) = scaleR r (f x)" and "\x. norm (f x) \ norm x * K" shows "bounded_linear f" by standard (blast intro: assms)+ locale bounded_bilinear = fixes prod :: "'a::real_normed_vector \ 'b::real_normed_vector \ 'c::real_normed_vector" (infixl "**" 70) assumes add_left: "prod (a + a') b = prod a b + prod a' b" and add_right: "prod a (b + b') = prod a b + prod a b'" and scaleR_left: "prod (scaleR r a) b = scaleR r (prod a b)" and scaleR_right: "prod a (scaleR r b) = scaleR r (prod a b)" and bounded: "\K. \a b. norm (prod a b) \ norm a * norm b * K" begin lemma pos_bounded: "\K>0. \a b. norm (a ** b) \ norm a * norm b * K" proof - obtain K where "\a b. norm (a ** b) \ norm a * norm b * K" using bounded by blast then have "norm (a ** b) \ norm a * norm b * (max 1 K)" for a b by (rule order.trans) (simp add: mult_left_mono) then show ?thesis by force qed lemma nonneg_bounded: "\K\0. \a b. norm (a ** b) \ norm a * norm b * K" using pos_bounded by (auto intro: order_less_imp_le) lemma additive_right: "additive (\b. prod a b)" by (rule additive.intro, rule add_right) lemma additive_left: "additive (\a. prod a b)" by (rule additive.intro, rule add_left) lemma zero_left: "prod 0 b = 0" by (rule additive.zero [OF additive_left]) lemma zero_right: "prod a 0 = 0" by (rule additive.zero [OF additive_right]) lemma minus_left: "prod (- a) b = - prod a b" by (rule additive.minus [OF additive_left]) lemma minus_right: "prod a (- b) = - prod a b" by (rule additive.minus [OF additive_right]) lemma diff_left: "prod (a - a') b = prod a b - prod a' b" by (rule additive.diff [OF additive_left]) lemma diff_right: "prod a (b - b') = prod a b - prod a b'" by (rule additive.diff [OF additive_right]) lemma sum_left: "prod (sum g S) x = sum ((\i. prod (g i) x)) S" by (rule additive.sum [OF additive_left]) lemma sum_right: "prod x (sum g S) = sum ((\i. (prod x (g i)))) S" by (rule additive.sum [OF additive_right]) lemma bounded_linear_left: "bounded_linear (\a. a ** b)" proof - obtain K where "\a b. norm (a ** b) \ norm a * norm b * K" using pos_bounded by blast then show ?thesis by (rule_tac K="norm b * K" in bounded_linear_intro) (auto simp: algebra_simps scaleR_left add_left) qed lemma bounded_linear_right: "bounded_linear (\b. a ** b)" proof - obtain K where "\a b. norm (a ** b) \ norm a * norm b * K" using pos_bounded by blast then show ?thesis by (rule_tac K="norm a * K" in bounded_linear_intro) (auto simp: algebra_simps scaleR_right add_right) qed lemma prod_diff_prod: "(x ** y - a ** b) = (x - a) ** (y - b) + (x - a) ** b + a ** (y - b)" by (simp add: diff_left diff_right) lemma flip: "bounded_bilinear (\x y. y ** x)" proof show "\K. \a b. norm (b ** a) \ norm a * norm b * K" by (metis bounded mult.commute) qed (simp_all add: add_right add_left scaleR_right scaleR_left) lemma comp1: assumes "bounded_linear g" shows "bounded_bilinear (\x. (**) (g x))" proof unfold_locales interpret g: bounded_linear g by fact show "\a a' b. g (a + a') ** b = g a ** b + g a' ** b" "\a b b'. g a ** (b + b') = g a ** b + g a ** b'" "\r a b. g (r *\<^sub>R a) ** b = r *\<^sub>R (g a ** b)" "\a r b. g a ** (r *\<^sub>R b) = r *\<^sub>R (g a ** b)" by (auto simp: g.add add_left add_right g.scaleR scaleR_left scaleR_right) from g.nonneg_bounded nonneg_bounded obtain K L where nn: "0 \ K" "0 \ L" and K: "\x. norm (g x) \ norm x * K" and L: "\a b. norm (a ** b) \ norm a * norm b * L" by auto have "norm (g a ** b) \ norm a * K * norm b * L" for a b by (auto intro!: order_trans[OF K] order_trans[OF L] mult_mono simp: nn) then show "\K. \a b. norm (g a ** b) \ norm a * norm b * K" by (auto intro!: exI[where x="K * L"] simp: ac_simps) qed lemma comp: "bounded_linear f \ bounded_linear g \ bounded_bilinear (\x y. f x ** g y)" by (rule bounded_bilinear.flip[OF bounded_bilinear.comp1[OF bounded_bilinear.flip[OF comp1]]]) end lemma bounded_linear_ident[simp]: "bounded_linear (\x. x)" by standard (auto intro!: exI[of _ 1]) lemma bounded_linear_zero[simp]: "bounded_linear (\x. 0)" by standard (auto intro!: exI[of _ 1]) lemma bounded_linear_add: assumes "bounded_linear f" and "bounded_linear g" shows "bounded_linear (\x. f x + g x)" proof - interpret f: bounded_linear f by fact interpret g: bounded_linear g by fact show ?thesis proof from f.bounded obtain Kf where Kf: "norm (f x) \ norm x * Kf" for x by blast from g.bounded obtain Kg where Kg: "norm (g x) \ norm x * Kg" for x by blast show "\K. \x. norm (f x + g x) \ norm x * K" using add_mono[OF Kf Kg] by (intro exI[of _ "Kf + Kg"]) (auto simp: field_simps intro: norm_triangle_ineq order_trans) qed (simp_all add: f.add g.add f.scaleR g.scaleR scaleR_right_distrib) qed lemma bounded_linear_minus: assumes "bounded_linear f" shows "bounded_linear (\x. - f x)" proof - interpret f: bounded_linear f by fact show ?thesis by unfold_locales (simp_all add: f.add f.scaleR f.bounded) qed lemma bounded_linear_sub: "bounded_linear f \ bounded_linear g \ bounded_linear (\x. f x - g x)" using bounded_linear_add[of f "\x. - g x"] bounded_linear_minus[of g] by (auto simp: algebra_simps) lemma bounded_linear_sum: fixes f :: "'i \ 'a::real_normed_vector \ 'b::real_normed_vector" shows "(\i. i \ I \ bounded_linear (f i)) \ bounded_linear (\x. \i\I. f i x)" by (induct I rule: infinite_finite_induct) (auto intro!: bounded_linear_add) lemma bounded_linear_compose: assumes "bounded_linear f" and "bounded_linear g" shows "bounded_linear (\x. f (g x))" proof - interpret f: bounded_linear f by fact interpret g: bounded_linear g by fact show ?thesis proof unfold_locales show "f (g (x + y)) = f (g x) + f (g y)" for x y by (simp only: f.add g.add) show "f (g (scaleR r x)) = scaleR r (f (g x))" for r x by (simp only: f.scaleR g.scaleR) from f.pos_bounded obtain Kf where f: "\x. norm (f x) \ norm x * Kf" and Kf: "0 < Kf" by blast from g.pos_bounded obtain Kg where g: "\x. norm (g x) \ norm x * Kg" by blast show "\K. \x. norm (f (g x)) \ norm x * K" proof (intro exI allI) fix x have "norm (f (g x)) \ norm (g x) * Kf" using f . also have "\ \ (norm x * Kg) * Kf" using g Kf [THEN order_less_imp_le] by (rule mult_right_mono) also have "(norm x * Kg) * Kf = norm x * (Kg * Kf)" by (rule mult.assoc) finally show "norm (f (g x)) \ norm x * (Kg * Kf)" . qed qed qed lemma bounded_bilinear_mult: "bounded_bilinear ((*) :: 'a \ 'a \ 'a::real_normed_algebra)" proof (rule bounded_bilinear.intro) show "\K. \a b::'a. norm (a * b) \ norm a * norm b * K" by (rule_tac x=1 in exI) (simp add: norm_mult_ineq) qed (auto simp: algebra_simps) lemma bounded_linear_mult_left: "bounded_linear (\x::'a::real_normed_algebra. x * y)" using bounded_bilinear_mult by (rule bounded_bilinear.bounded_linear_left) lemma bounded_linear_mult_right: "bounded_linear (\y::'a::real_normed_algebra. x * y)" using bounded_bilinear_mult by (rule bounded_bilinear.bounded_linear_right) lemmas bounded_linear_mult_const = bounded_linear_mult_left [THEN bounded_linear_compose] lemmas bounded_linear_const_mult = bounded_linear_mult_right [THEN bounded_linear_compose] lemma bounded_linear_divide: "bounded_linear (\x. x / y)" for y :: "'a::real_normed_field" unfolding divide_inverse by (rule bounded_linear_mult_left) lemma bounded_bilinear_scaleR: "bounded_bilinear scaleR" proof (rule bounded_bilinear.intro) show "\K. \a b. norm (a *\<^sub>R b) \ norm a * norm b * K" using less_eq_real_def by auto qed (auto simp: algebra_simps) lemma bounded_linear_scaleR_left: "bounded_linear (\r. scaleR r x)" using bounded_bilinear_scaleR by (rule bounded_bilinear.bounded_linear_left) lemma bounded_linear_scaleR_right: "bounded_linear (\x. scaleR r x)" using bounded_bilinear_scaleR by (rule bounded_bilinear.bounded_linear_right) lemmas bounded_linear_scaleR_const = bounded_linear_scaleR_left[THEN bounded_linear_compose] lemmas bounded_linear_const_scaleR = bounded_linear_scaleR_right[THEN bounded_linear_compose] lemma bounded_linear_of_real: "bounded_linear (\r. of_real r)" unfolding of_real_def by (rule bounded_linear_scaleR_left) lemma real_bounded_linear: "bounded_linear f \ (\c::real. f = (\x. x * c))" for f :: "real \ real" proof - { fix x assume "bounded_linear f" then interpret bounded_linear f . from scaleR[of x 1] have "f x = x * f 1" by simp } then show ?thesis by (auto intro: exI[of _ "f 1"] bounded_linear_mult_left) qed instance real_normed_algebra_1 \ perfect_space proof fix x::'a have "\e. 0 < e \ \y. norm (y - x) < e \ y \ x" by (rule_tac x = "x + of_real (e/2)" in exI) auto then show "\ open {x}" by (clarsimp simp: open_dist dist_norm) qed subsection \Filters and Limits on Metric Space\ lemma (in metric_space) nhds_metric: "nhds x = (INF e\{0 <..}. principal {y. dist y x < e})" unfolding nhds_def proof (safe intro!: INF_eq) fix S assume "open S" "x \ S" then obtain e where "{y. dist y x < e} \ S" "0 < e" by (auto simp: open_dist subset_eq) then show "\e\{0<..}. principal {y. dist y x < e} \ principal S" by auto qed (auto intro!: exI[of _ "{y. dist x y < e}" for e] open_ball simp: dist_commute) lemma (in metric_space) tendsto_iff: "(f \ l) F \ (\e>0. eventually (\x. dist (f x) l < e) F)" unfolding nhds_metric filterlim_INF filterlim_principal by auto lemma tendsto_dist_iff: "((f \ l) F) \ (((\x. dist (f x) l) \ 0) F)" unfolding tendsto_iff by simp lemma (in metric_space) tendstoI [intro?]: "(\e. 0 < e \ eventually (\x. dist (f x) l < e) F) \ (f \ l) F" by (auto simp: tendsto_iff) lemma (in metric_space) tendstoD: "(f \ l) F \ 0 < e \ eventually (\x. dist (f x) l < e) F" by (auto simp: tendsto_iff) lemma (in metric_space) eventually_nhds_metric: "eventually P (nhds a) \ (\d>0. \x. dist x a < d \ P x)" unfolding nhds_metric by (subst eventually_INF_base) (auto simp: eventually_principal Bex_def subset_eq intro: exI[of _ "min a b" for a b]) lemma eventually_at: "eventually P (at a within S) \ (\d>0. \x\S. x \ a \ dist x a < d \ P x)" for a :: "'a :: metric_space" by (auto simp: eventually_at_filter eventually_nhds_metric) lemma frequently_at: "frequently P (at a within S) \ (\d>0. \x\S. x \ a \ dist x a < d \ P x)" for a :: "'a :: metric_space" unfolding frequently_def eventually_at by auto lemma eventually_at_le: "eventually P (at a within S) \ (\d>0. \x\S. x \ a \ dist x a \ d \ P x)" for a :: "'a::metric_space" unfolding eventually_at_filter eventually_nhds_metric apply safe apply (rule_tac x="d / 2" in exI, auto) done lemma eventually_at_left_real: "a > (b :: real) \ eventually (\x. x \ {b<.. eventually (\x. x \ {a<.. a) F" and le: "eventually (\x. dist (g x) b \ dist (f x) a) F" shows "(g \ b) F" proof (rule tendstoI) fix e :: real assume "0 < e" with f have "eventually (\x. dist (f x) a < e) F" by (rule tendstoD) with le show "eventually (\x. dist (g x) b < e) F" using le_less_trans by (rule eventually_elim2) qed lemma filterlim_real_sequentially: "LIM x sequentially. real x :> at_top" proof (clarsimp simp: filterlim_at_top) fix Z show "\\<^sub>F x in sequentially. Z \ real x" by (meson eventually_sequentiallyI nat_ceiling_le_eq) qed lemma filterlim_nat_sequentially: "filterlim nat sequentially at_top" proof - have "\\<^sub>F x in at_top. Z \ nat x" for Z by (auto intro!: eventually_at_top_linorderI[where c="int Z"]) then show ?thesis unfolding filterlim_at_top .. qed lemma filterlim_floor_sequentially: "filterlim floor at_top at_top" proof - have "\\<^sub>F x in at_top. Z \ \x\" for Z by (auto simp: le_floor_iff intro!: eventually_at_top_linorderI[where c="of_int Z"]) then show ?thesis unfolding filterlim_at_top .. qed lemma filterlim_sequentially_iff_filterlim_real: "filterlim f sequentially F \ filterlim (\x. real (f x)) at_top F" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using filterlim_compose filterlim_real_sequentially by blast next assume R: ?rhs show ?lhs proof - have "filterlim (\x. nat (floor (real (f x)))) sequentially F" by (intro filterlim_compose[OF filterlim_nat_sequentially] filterlim_compose[OF filterlim_floor_sequentially] R) then show ?thesis by simp qed qed subsubsection \Limits of Sequences\ lemma lim_sequentially: "X \ L \ (\r>0. \no. \n\no. dist (X n) L < r)" for L :: "'a::metric_space" unfolding tendsto_iff eventually_sequentially .. lemmas LIMSEQ_def = lim_sequentially (*legacy binding*) lemma LIMSEQ_iff_nz: "X \ L \ (\r>0. \no>0. \n\no. dist (X n) L < r)" for L :: "'a::metric_space" unfolding lim_sequentially by (metis Suc_leD zero_less_Suc) lemma metric_LIMSEQ_I: "(\r. 0 < r \ \no. \n\no. dist (X n) L < r) \ X \ L" for L :: "'a::metric_space" by (simp add: lim_sequentially) lemma metric_LIMSEQ_D: "X \ L \ 0 < r \ \no. \n\no. dist (X n) L < r" for L :: "'a::metric_space" by (simp add: lim_sequentially) lemma LIMSEQ_norm_0: assumes "\n::nat. norm (f n) < 1 / real (Suc n)" shows "f \ 0" proof (rule metric_LIMSEQ_I) fix \ :: "real" assume "\ > 0" then obtain N::nat where "\ > inverse N" "N > 0" by (metis neq0_conv real_arch_inverse) then have "norm (f n) < \" if "n \ N" for n proof - have "1 / (Suc n) \ 1 / N" using \0 < N\ inverse_of_nat_le le_SucI that by blast also have "\ < \" by (metis (no_types) \inverse (real N) < \\ inverse_eq_divide) finally show ?thesis by (meson assms less_eq_real_def not_le order_trans) qed then show "\no. \n\no. dist (f n) 0 < \" by auto qed subsubsection \Limits of Functions\ lemma LIM_def: "f \a\ L \ (\r > 0. \s > 0. \x. x \ a \ dist x a < s \ dist (f x) L < r)" for a :: "'a::metric_space" and L :: "'b::metric_space" unfolding tendsto_iff eventually_at by simp lemma metric_LIM_I: "(\r. 0 < r \ \s>0. \x. x \ a \ dist x a < s \ dist (f x) L < r) \ f \a\ L" for a :: "'a::metric_space" and L :: "'b::metric_space" by (simp add: LIM_def) lemma metric_LIM_D: "f \a\ L \ 0 < r \ \s>0. \x. x \ a \ dist x a < s \ dist (f x) L < r" for a :: "'a::metric_space" and L :: "'b::metric_space" by (simp add: LIM_def) lemma metric_LIM_imp_LIM: fixes l :: "'a::metric_space" and m :: "'b::metric_space" assumes f: "f \a\ l" and le: "\x. x \ a \ dist (g x) m \ dist (f x) l" shows "g \a\ m" by (rule metric_tendsto_imp_tendsto [OF f]) (auto simp: eventually_at_topological le) lemma metric_LIM_equal2: fixes a :: "'a::metric_space" assumes "g \a\ l" "0 < R" and "\x. x \ a \ dist x a < R \ f x = g x" shows "f \a\ l" proof - have "\S. \open S; l \ S; \\<^sub>F x in at a. g x \ S\ \ \\<^sub>F x in at a. f x \ S" apply (simp add: eventually_at) by (metis assms(2) assms(3) dual_order.strict_trans linorder_neqE_linordered_idom) then show ?thesis using assms by (simp add: tendsto_def) qed lemma metric_LIM_compose2: fixes a :: "'a::metric_space" assumes f: "f \a\ b" and g: "g \b\ c" and inj: "\d>0. \x. x \ a \ dist x a < d \ f x \ b" shows "(\x. g (f x)) \a\ c" using inj by (intro tendsto_compose_eventually[OF g f]) (auto simp: eventually_at) lemma metric_isCont_LIM_compose2: fixes f :: "'a :: metric_space \ _" assumes f [unfolded isCont_def]: "isCont f a" and g: "g \f a\ l" and inj: "\d>0. \x. x \ a \ dist x a < d \ f x \ f a" shows "(\x. g (f x)) \a\ l" by (rule metric_LIM_compose2 [OF f g inj]) subsection \Complete metric spaces\ subsection \Cauchy sequences\ lemma (in metric_space) Cauchy_def: "Cauchy X = (\e>0. \M. \m\M. \n\M. dist (X m) (X n) < e)" proof - have *: "eventually P (INF M. principal {(X m, X n) | n m. m \ M \ n \ M}) \ (\M. \m\M. \n\M. P (X m, X n))" for P apply (subst eventually_INF_base) subgoal by simp subgoal for a b by (intro bexI[of _ "max a b"]) (auto simp: eventually_principal subset_eq) subgoal by (auto simp: eventually_principal, blast) done have "Cauchy X \ (INF M. principal {(X m, X n) | n m. m \ M \ n \ M}) \ uniformity" unfolding Cauchy_uniform_iff le_filter_def * .. also have "\ = (\e>0. \M. \m\M. \n\M. dist (X m) (X n) < e)" unfolding uniformity_dist le_INF_iff by (auto simp: * le_principal) finally show ?thesis . qed lemma (in metric_space) Cauchy_altdef: "Cauchy f \ (\e>0. \M. \m\M. \n>m. dist (f m) (f n) < e)" (is "?lhs \ ?rhs") proof assume ?rhs show ?lhs unfolding Cauchy_def proof (intro allI impI) fix e :: real assume e: "e > 0" with \?rhs\ obtain M where M: "m \ M \ n > m \ dist (f m) (f n) < e" for m n by blast have "dist (f m) (f n) < e" if "m \ M" "n \ M" for m n using M[of m n] M[of n m] e that by (cases m n rule: linorder_cases) (auto simp: dist_commute) then show "\M. \m\M. \n\M. dist (f m) (f n) < e" by blast qed next assume ?lhs show ?rhs proof (intro allI impI) fix e :: real assume e: "e > 0" with \Cauchy f\ obtain M where "\m n. m \ M \ n \ M \ dist (f m) (f n) < e" unfolding Cauchy_def by blast then show "\M. \m\M. \n>m. dist (f m) (f n) < e" by (intro exI[of _ M]) force qed qed lemma (in metric_space) Cauchy_altdef2: "Cauchy s \ (\e>0. \N::nat. \n\N. dist(s n)(s N) < e)" (is "?lhs = ?rhs") proof assume "Cauchy s" then show ?rhs by (force simp: Cauchy_def) next assume ?rhs { fix e::real assume "e>0" with \?rhs\ obtain N where N: "\n\N. dist (s n) (s N) < e/2" by (erule_tac x="e/2" in allE) auto { fix n m assume nm: "N \ m \ N \ n" then have "dist (s m) (s n) < e" using N using dist_triangle_half_l[of "s m" "s N" "e" "s n"] by blast } then have "\N. \m n. N \ m \ N \ n \ dist (s m) (s n) < e" by blast } then have ?lhs unfolding Cauchy_def by blast then show ?lhs by blast qed lemma (in metric_space) metric_CauchyI: "(\e. 0 < e \ \M. \m\M. \n\M. dist (X m) (X n) < e) \ Cauchy X" by (simp add: Cauchy_def) lemma (in metric_space) CauchyI': "(\e. 0 < e \ \M. \m\M. \n>m. dist (X m) (X n) < e) \ Cauchy X" unfolding Cauchy_altdef by blast lemma (in metric_space) metric_CauchyD: "Cauchy X \ 0 < e \ \M. \m\M. \n\M. dist (X m) (X n) < e" by (simp add: Cauchy_def) lemma (in metric_space) metric_Cauchy_iff2: "Cauchy X = (\j. (\M. \m \ M. \n \ M. dist (X m) (X n) < inverse(real (Suc j))))" apply (auto simp add: Cauchy_def) by (metis less_trans of_nat_Suc reals_Archimedean) lemma Cauchy_iff2: "Cauchy X \ (\j. (\M. \m \ M. \n \ M. \X m - X n\ < inverse (real (Suc j))))" by (simp only: metric_Cauchy_iff2 dist_real_def) lemma lim_1_over_n [tendsto_intros]: "((\n. 1 / of_nat n) \ (0::'a::real_normed_field)) sequentially" proof (subst lim_sequentially, intro allI impI exI) fix e::real and n assume e: "e > 0" have "inverse e < of_nat (nat \inverse e + 1\)" by linarith also assume "n \ nat \inverse e + 1\" finally show "dist (1 / of_nat n :: 'a) 0 < e" using e by (simp add: field_split_simps norm_divide) qed lemma (in metric_space) complete_def: shows "complete S = (\f. (\n. f n \ S) \ Cauchy f \ (\l\S. f \ l))" unfolding complete_uniform proof safe fix f :: "nat \ 'a" assume f: "\n. f n \ S" "Cauchy f" and *: "\F\principal S. F \ bot \ cauchy_filter F \ (\x\S. F \ nhds x)" then show "\l\S. f \ l" unfolding filterlim_def using f by (intro *[rule_format]) (auto simp: filtermap_sequentually_ne_bot le_principal eventually_filtermap Cauchy_uniform) next fix F :: "'a filter" assume "F \ principal S" "F \ bot" "cauchy_filter F" assume seq: "\f. (\n. f n \ S) \ Cauchy f \ (\l\S. f \ l)" from \F \ principal S\ \cauchy_filter F\ have FF_le: "F \\<^sub>F F \ uniformity_on S" by (simp add: cauchy_filter_def principal_prod_principal[symmetric] prod_filter_mono) let ?P = "\P e. eventually P F \ (\x. P x \ x \ S) \ (\x y. P x \ P y \ dist x y < e)" have P: "\P. ?P P \" if "0 < \" for \ :: real proof - from that have "eventually (\(x, y). x \ S \ y \ S \ dist x y < \) (uniformity_on S)" by (auto simp: eventually_inf_principal eventually_uniformity_metric) from filter_leD[OF FF_le this] show ?thesis by (auto simp: eventually_prod_same) qed have "\P. \n. ?P (P n) (1 / Suc n) \ P (Suc n) \ P n" proof (rule dependent_nat_choice) show "\P. ?P P (1 / Suc 0)" using P[of 1] by auto next fix P n assume "?P P (1/Suc n)" moreover obtain Q where "?P Q (1 / Suc (Suc n))" using P[of "1/Suc (Suc n)"] by auto ultimately show "\Q. ?P Q (1 / Suc (Suc n)) \ Q \ P" by (intro exI[of _ "\x. P x \ Q x"]) (auto simp: eventually_conj_iff) qed then obtain P where P: "eventually (P n) F" "P n x \ x \ S" "P n x \ P n y \ dist x y < 1 / Suc n" "P (Suc n) \ P n" for n x y by metis have "antimono P" using P(4) unfolding decseq_Suc_iff le_fun_def by blast obtain X where X: "P n (X n)" for n using P(1)[THEN eventually_happens'[OF \F \ bot\]] by metis have "Cauchy X" unfolding metric_Cauchy_iff2 inverse_eq_divide proof (intro exI allI impI) fix j m n :: nat assume "j \ m" "j \ n" with \antimono P\ X have "P j (X m)" "P j (X n)" by (auto simp: antimono_def) then show "dist (X m) (X n) < 1 / Suc j" by (rule P) qed moreover have "\n. X n \ S" using P(2) X by auto ultimately obtain x where "X \ x" "x \ S" using seq by blast show "\x\S. F \ nhds x" proof (rule bexI) have "eventually (\y. dist y x < e) F" if "0 < e" for e :: real proof - from that have "(\n. 1 / Suc n :: real) \ 0 \ 0 < e / 2" by (subst filterlim_sequentially_Suc) (auto intro!: lim_1_over_n) then have "\\<^sub>F n in sequentially. dist (X n) x < e / 2 \ 1 / Suc n < e / 2" using \X \ x\ unfolding tendsto_iff order_tendsto_iff[where 'a=real] eventually_conj_iff by blast then obtain n where "dist x (X n) < e / 2" "1 / Suc n < e / 2" by (auto simp: eventually_sequentially dist_commute) show ?thesis using \eventually (P n) F\ proof eventually_elim case (elim y) then have "dist y (X n) < 1 / Suc n" by (intro X P) also have "\ < e / 2" by fact finally show "dist y x < e" by (rule dist_triangle_half_l) fact qed qed then show "F \ nhds x" unfolding nhds_metric le_INF_iff le_principal by auto qed fact qed text\apparently unused\ lemma (in metric_space) totally_bounded_metric: "totally_bounded S \ (\e>0. \k. finite k \ S \ (\x\k. {y. dist x y < e}))" unfolding totally_bounded_def eventually_uniformity_metric imp_ex apply (subst all_comm) apply (intro arg_cong[where f=All] ext, safe) subgoal for e apply (erule allE[of _ "\(x, y). dist x y < e"]) apply auto done subgoal for e P k apply (intro exI[of _ k]) apply (force simp: subset_eq) done done subsubsection \Cauchy Sequences are Convergent\ (* TODO: update to uniform_space *) class complete_space = metric_space + assumes Cauchy_convergent: "Cauchy X \ convergent X" lemma Cauchy_convergent_iff: "Cauchy X \ convergent X" for X :: "nat \ 'a::complete_space" by (blast intro: Cauchy_convergent convergent_Cauchy) text \To prove that a Cauchy sequence converges, it suffices to show that a subsequence converges.\ lemma Cauchy_converges_subseq: fixes u::"nat \ 'a::metric_space" assumes "Cauchy u" "strict_mono r" "(u \ r) \ l" shows "u \ l" proof - have *: "eventually (\n. dist (u n) l < e) sequentially" if "e > 0" for e proof - have "e/2 > 0" using that by auto then obtain N1 where N1: "\m n. m \ N1 \ n \ N1 \ dist (u m) (u n) < e/2" using \Cauchy u\ unfolding Cauchy_def by blast obtain N2 where N2: "\n. n \ N2 \ dist ((u \ r) n) l < e / 2" using order_tendstoD(2)[OF iffD1[OF tendsto_dist_iff \(u \ r) \ l\] \e/2 > 0\] unfolding eventually_sequentially by auto have "dist (u n) l < e" if "n \ max N1 N2" for n proof - have "dist (u n) l \ dist (u n) ((u \ r) n) + dist ((u \ r) n) l" by (rule dist_triangle) also have "\ < e/2 + e/2" proof (intro add_strict_mono) show "dist (u n) ((u \ r) n) < e / 2" using N1[of n "r n"] N2[of n] that unfolding comp_def by (meson assms(2) le_trans max.bounded_iff strict_mono_imp_increasing) show "dist ((u \ r) n) l < e / 2" using N2 that by auto qed finally show ?thesis by simp qed then show ?thesis unfolding eventually_sequentially by blast qed have "(\n. dist (u n) l) \ 0" by (simp add: less_le_trans * order_tendstoI) then show ?thesis using tendsto_dist_iff by auto qed subsection \The set of real numbers is a complete metric space\ text \ Proof that Cauchy sequences converge based on the one from \<^url>\http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html\ \ text \ If sequence \<^term>\X\ is Cauchy, then its limit is the lub of \<^term>\{r::real. \N. \n\N. r < X n}\ \ lemma increasing_LIMSEQ: fixes f :: "nat \ real" assumes inc: "\n. f n \ f (Suc n)" and bdd: "\n. f n \ l" and en: "\e. 0 < e \ \n. l \ f n + e" shows "f \ l" proof (rule increasing_tendsto) fix x assume "x < l" with dense[of 0 "l - x"] obtain e where "0 < e" "e < l - x" by auto from en[OF \0 < e\] obtain n where "l - e \ f n" by (auto simp: field_simps) with \e < l - x\ \0 < e\ have "x < f n" by simp with incseq_SucI[of f, OF inc] show "eventually (\n. x < f n) sequentially" by (auto simp: eventually_sequentially incseq_def intro: less_le_trans) qed (use bdd in auto) lemma real_Cauchy_convergent: fixes X :: "nat \ real" assumes X: "Cauchy X" shows "convergent X" proof - define S :: "real set" where "S = {x. \N. \n\N. x < X n}" then have mem_S: "\N x. \n\N. x < X n \ x \ S" by auto have bound_isUb: "y \ x" if N: "\n\N. X n < x" and "y \ S" for N and x y :: real proof - from that have "\M. \n\M. y < X n" by (simp add: S_def) then obtain M where "\n\M. y < X n" .. then have "y < X (max M N)" by simp also have "\ < x" using N by simp finally show ?thesis by (rule order_less_imp_le) qed obtain N where "\m\N. \n\N. dist (X m) (X n) < 1" using X[THEN metric_CauchyD, OF zero_less_one] by auto then have N: "\n\N. dist (X n) (X N) < 1" by simp have [simp]: "S \ {}" proof (intro exI ex_in_conv[THEN iffD1]) from N have "\n\N. X N - 1 < X n" by (simp add: abs_diff_less_iff dist_real_def) then show "X N - 1 \ S" by (rule mem_S) qed have [simp]: "bdd_above S" proof from N have "\n\N. X n < X N + 1" by (simp add: abs_diff_less_iff dist_real_def) then show "\s. s \ S \ s \ X N + 1" by (rule bound_isUb) qed have "X \ Sup S" proof (rule metric_LIMSEQ_I) fix r :: real assume "0 < r" then have r: "0 < r/2" by simp obtain N where "\n\N. \m\N. dist (X n) (X m) < r/2" using metric_CauchyD [OF X r] by auto then have "\n\N. dist (X n) (X N) < r/2" by simp then have N: "\n\N. X N - r/2 < X n \ X n < X N + r/2" by (simp only: dist_real_def abs_diff_less_iff) from N have "\n\N. X N - r/2 < X n" by blast then have "X N - r/2 \ S" by (rule mem_S) then have 1: "X N - r/2 \ Sup S" by (simp add: cSup_upper) from N have "\n\N. X n < X N + r/2" by blast from bound_isUb[OF this] have 2: "Sup S \ X N + r/2" by (intro cSup_least) simp_all show "\N. \n\N. dist (X n) (Sup S) < r" proof (intro exI allI impI) fix n assume n: "N \ n" from N n have "X n < X N + r/2" and "X N - r/2 < X n" by simp_all then show "dist (X n) (Sup S) < r" using 1 2 by (simp add: abs_diff_less_iff dist_real_def) qed qed then show ?thesis by (auto simp: convergent_def) qed instance real :: complete_space by intro_classes (rule real_Cauchy_convergent) class banach = real_normed_vector + complete_space instance real :: banach .. lemma tendsto_at_topI_sequentially: fixes f :: "real \ 'b::first_countable_topology" assumes *: "\X. filterlim X at_top sequentially \ (\n. f (X n)) \ y" shows "(f \ y) at_top" proof - obtain A where A: "decseq A" "open (A n)" "y \ A n" "nhds y = (INF n. principal (A n))" for n by (rule nhds_countable[of y]) (rule that) have "\m. \k. \x\k. f x \ A m" proof (rule ccontr) assume "\ (\m. \k. \x\k. f x \ A m)" then obtain m where "\k. \x\k. f x \ A m" by auto then have "\X. \n. (f (X n) \ A m) \ max n (X n) + 1 \ X (Suc n)" by (intro dependent_nat_choice) (auto simp del: max.bounded_iff) then obtain X where X: "\n. f (X n) \ A m" "\n. max n (X n) + 1 \ X (Suc n)" by auto have "1 \ n \ real n \ X n" for n using X[of "n - 1"] by auto then have "filterlim X at_top sequentially" by (force intro!: filterlim_at_top_mono[OF filterlim_real_sequentially] simp: eventually_sequentially) from topological_tendstoD[OF *[OF this] A(2, 3), of m] X(1) show False by auto qed then obtain k where "k m \ x \ f x \ A m" for m x by metis then show ?thesis unfolding at_top_def A by (intro filterlim_base[where i=k]) auto qed lemma tendsto_at_topI_sequentially_real: fixes f :: "real \ real" assumes mono: "mono f" and limseq: "(\n. f (real n)) \ y" shows "(f \ y) at_top" proof (rule tendstoI) fix e :: real assume "0 < e" with limseq obtain N :: nat where N: "N \ n \ \f (real n) - y\ < e" for n by (auto simp: lim_sequentially dist_real_def) have le: "f x \ y" for x :: real proof - obtain n where "x \ real_of_nat n" using real_arch_simple[of x] .. note monoD[OF mono this] also have "f (real_of_nat n) \ y" by (rule LIMSEQ_le_const[OF limseq]) (auto intro!: exI[of _ n] monoD[OF mono]) finally show ?thesis . qed have "eventually (\x. real N \ x) at_top" by (rule eventually_ge_at_top) then show "eventually (\x. dist (f x) y < e) at_top" proof eventually_elim case (elim x) with N[of N] le have "y - f (real N) < e" by auto moreover note monoD[OF mono elim] ultimately show "dist (f x) y < e" using le[of x] by (auto simp: dist_real_def field_simps) qed qed end