# HG changeset patch # User haftmann # Date 1665065799 0 # Node ID fbde7d1211fcc68942812a0f894fce24d75f60e6 # Parent 63bcec915790d4c4df505e0a76198151f7330534 euclidean division on gaussian numbers diff -r 63bcec915790 -r fbde7d1211fc src/HOL/Examples/Gauss_Numbers.thy --- a/src/HOL/Examples/Gauss_Numbers.thy Thu Oct 06 13:41:59 2022 +0000 +++ b/src/HOL/Examples/Gauss_Numbers.thy Thu Oct 06 14:16:39 2022 +0000 @@ -325,16 +325,81 @@ show \x div 0 = 0\ by (simp add: gauss_eq_iff) show \x * y div y = x\ if \y \ 0\ - using that - apply (auto simp add: gauss_eq_iff algebra_simps power2_eq_square) - apply (simp_all only: flip: mult.assoc distrib_right) - apply (simp_all add: mult.assoc [of _ \Re y\ \Re y\] mult.assoc [of _ \Im y\ \Im y\] mult.commute [of _ \Re x\] flip: distrib_left) - apply (simp_all add: sum_squares_eq_zero_iff mult.commute nonzero_mult_rdiv_cancel_right flip: distrib_left) - done + proof - + define Y where \Y = (Re y)\<^sup>2 + (Im y)\<^sup>2\ + moreover have \Y > 0\ + using that by (simp add: gauss_eq_0 less_le Y_def) + have *: \Im y * (Im y * Re x) + Re x * (Re y * Re y) = Re x * Y\ + \Im x * (Im y * Im y) + Im x * (Re y * Re y) = Im x * Y\ + \(Im y)\<^sup>2 + (Re y)\<^sup>2 = Y\ + by (simp_all add: power2_eq_square algebra_simps Y_def) + from \Y > 0\ show ?thesis + by (simp add: gauss_eq_iff algebra_simps) (simp add: * nonzero_mult_rdiv_cancel_right) + qed show \x div y * y + x mod y = x\ by (simp add: gauss_eq_iff) qed end +instantiation gauss :: euclidean_ring +begin + +definition euclidean_size_gauss :: \gauss \ nat\ + where \euclidean_size x = nat ((Re x)\<^sup>2 + (Im x)\<^sup>2)\ + +instance proof + show \euclidean_size (0::gauss) = 0\ + by (simp add: euclidean_size_gauss_def) + show \euclidean_size (x mod y) < euclidean_size y\ if \y \ 0\ for x y :: gauss + proof- + define X and Y and R and I + where \X = (Re x)\<^sup>2 + (Im x)\<^sup>2\ and \Y = (Re y)\<^sup>2 + (Im y)\<^sup>2\ + and \R = Re x * Re y + Im x * Im y\ and \I = Im x * Re y - Re x * Im y\ + with that have \0 < Y\ and rhs: \int (euclidean_size y) = Y\ + by (simp_all add: gauss_neq_0 euclidean_size_gauss_def) + have \X * Y = R\<^sup>2 + I\<^sup>2\ + by (simp add: R_def I_def X_def Y_def power2_eq_square algebra_simps) + let ?lhs = \X - I * (I rdiv Y) - R * (R rdiv Y) + - I rdiv Y * (I rmod Y) - R rdiv Y * (R rmod Y)\ + have \?lhs = X + Y * (R rdiv Y) * (R rdiv Y) + Y * (I rdiv Y) * (I rdiv Y) + - 2 * (R rdiv Y * R + I rdiv Y * I)\ + by (simp flip: minus_rmod_eq_mult_rdiv add: algebra_simps) + also have \\ = (Re (x mod y))\<^sup>2 + (Im (x mod y))\<^sup>2\ + by (simp add: X_def Y_def R_def I_def algebra_simps power2_eq_square) + finally have lhs: \int (euclidean_size (x mod y)) = ?lhs\ + by (simp add: euclidean_size_gauss_def) + have \?lhs * Y = (I rmod Y)\<^sup>2 + (R rmod Y)\<^sup>2\ + apply (simp add: algebra_simps power2_eq_square \X * Y = R\<^sup>2 + I\<^sup>2\) + apply (simp flip: mult.assoc add.assoc minus_rmod_eq_mult_rdiv) + apply (simp add: algebra_simps) + done + also have \\ \ (Y div 2)\<^sup>2 + (Y div 2)\<^sup>2\ + by (rule add_mono) (use \Y > 0\ abs_rmod_less_equal [of Y] in \simp_all add: power2_le_iff_abs_le\) + also have \\ < Y\<^sup>2\ + using \Y > 0\ by (cases \Y = 1\) (simp_all add: power2_eq_square mult_le_less_imp_less flip: mult.assoc) + finally have \?lhs * Y < Y\<^sup>2\ . + with \Y > 0\ have \?lhs < Y\ + by (simp add: power2_eq_square) + then have \int (euclidean_size (x mod y)) < int (euclidean_size y)\ + by (simp only: lhs rhs) + then show ?thesis + by simp + qed + show \euclidean_size x \ euclidean_size (x * y)\ if \y \ 0\ for x y :: gauss + proof - + from that have \euclidean_size y > 0\ + by (simp add: euclidean_size_gauss_def gauss_neq_0) + then have \euclidean_size x \ euclidean_size x * euclidean_size y\ + by simp + also have \\ = nat (((Re x)\<^sup>2 + (Im x)\<^sup>2) * ((Re y)\<^sup>2 + (Im y)\<^sup>2))\ + by (simp add: euclidean_size_gauss_def nat_mult_distrib) + also have \\ = euclidean_size (x * y)\ + by (simp add: euclidean_size_gauss_def eq_nat_nat_iff) (simp add: algebra_simps power2_eq_square) + finally show ?thesis . + qed +qed + end + +end diff -r 63bcec915790 -r fbde7d1211fc src/HOL/Library/Rounded_Division.thy --- a/src/HOL/Library/Rounded_Division.thy Thu Oct 06 13:41:59 2022 +0000 +++ b/src/HOL/Library/Rounded_Division.thy Thu Oct 06 14:16:39 2022 +0000 @@ -7,6 +7,10 @@ imports Main begin +lemma off_iff_abs_mod_2_eq_one: + \odd l \ \l\ mod 2 = 1\ for l :: int + by (simp flip: odd_iff_mod_2_eq_one) + definition rounded_divide :: \int \ int \ int\ (infixl \rdiv\ 70) where \k rdiv l = sgn l * ((k + \l\ div 2) div \l\)\ @@ -119,4 +123,39 @@ by (simp only: rounded_modulo_def power_abs abs_numeral flip: take_bit_eq_mod) (simp add: signed_take_bit_eq_take_bit_shift) +lemma rmod_less_divisor: + \k rmod l < \l\ - \l\ div 2\ if \l \ 0\ + using that pos_mod_bound [of \\l\\] by (simp add: rounded_modulo_def) + +lemma rmod_less_equal_divisor: + \k rmod l \ \l\ div 2\ if \l \ 0\ +proof - + from that rmod_less_divisor [of l k] + have \k rmod l < \l\ - \l\ div 2\ + by simp + also have \\l\ - \l\ div 2 = \l\ div 2 + of_bool (odd l)\ + by auto + finally show ?thesis + by (cases \even l\) simp_all +qed + +lemma divisor_less_equal_rmod': + \\l\ div 2 - \l\ \ k rmod l\ if \l \ 0\ +proof - + have \0 \ (k + \l\ div 2) mod \l\\ + using that pos_mod_sign [of \\l\\] by simp + then show ?thesis + by (simp_all add: rounded_modulo_def) +qed + +lemma divisor_less_equal_rmod: + \- (\l\ div 2) \ k rmod l\ if \l \ 0\ + using that divisor_less_equal_rmod' [of l k] + by (simp add: rounded_modulo_def) + +lemma abs_rmod_less_equal: + \\k rmod l\ \ \l\ div 2\ if \l \ 0\ + using that divisor_less_equal_rmod [of l k] + by (simp add: abs_le_iff rmod_less_equal_divisor) + end