diff --git a/src/HOL/Examples/Gauss_Numbers.thy b/src/HOL/Examples/Gauss_Numbers.thy --- a/src/HOL/Examples/Gauss_Numbers.thy +++ b/src/HOL/Examples/Gauss_Numbers.thy @@ -1,329 +1,329 @@ -(* Author: Florian Haftmann, TU Muenchen; based on existing material on gauss numbers\ +(* Author: Florian Haftmann, TU Muenchen; based on existing material on complex numbers\ *) section \Gauss Numbers: integral gauss numbers\ theory Gauss_Numbers -imports Main + imports "HOL-Library.Rounded_Division" begin codatatype gauss = Gauss (Re: int) (Im: int) lemma gauss_eqI [intro?]: \x = y\ if \Re x = Re y\ \Im x = Im y\ by (rule gauss.expand) (use that in simp) lemma gauss_eq_iff: \x = y \ Re x = Re y \ Im x = Im y\ by (auto intro: gauss_eqI) subsection \Basic arithmetic\ instantiation gauss :: comm_ring_1 begin primcorec zero_gauss :: \gauss\ where \Re 0 = 0\ | \Im 0 = 0\ primcorec one_gauss :: \gauss\ where \Re 1 = 1\ | \Im 1 = 0\ primcorec plus_gauss :: \gauss \ gauss \ gauss\ where \Re (x + y) = Re x + Re y\ | \Im (x + y) = Im x + Im y\ primcorec uminus_gauss :: \gauss \ gauss\ where \Re (- x) = - Re x\ | \Im (- x) = - Im x\ primcorec minus_gauss :: \gauss \ gauss \ gauss\ where \Re (x - y) = Re x - Re y\ | \Im (x - y) = Im x - Im y\ primcorec times_gauss :: \gauss \ gauss \ gauss\ where \Re (x * y) = Re x * Re y - Im x * Im y\ | \Im (x * y) = Re x * Im y + Im x * Re y\ instance by standard (simp_all add: gauss_eq_iff algebra_simps) end lemma of_nat_gauss: \of_nat n = Gauss (int n) 0\ by (induction n) (simp_all add: gauss_eq_iff) lemma numeral_gauss: \numeral n = Gauss (numeral n) 0\ proof - have \numeral n = (of_nat (numeral n) :: gauss)\ by simp also have \\ = Gauss (of_nat (numeral n)) 0\ by (simp add: of_nat_gauss) finally show ?thesis by simp qed lemma of_int_gauss: \of_int k = Gauss k 0\ by (simp add: gauss_eq_iff of_int_of_nat of_nat_gauss) lemma conversion_simps [simp]: \Re (numeral m) = numeral m\ \Im (numeral m) = 0\ \Re (of_nat n) = int n\ \Im (of_nat n) = 0\ \Re (of_int k) = k\ \Im (of_int k) = 0\ by (simp_all add: numeral_gauss of_nat_gauss of_int_gauss) lemma gauss_eq_0: \z = 0 \ (Re z)\<^sup>2 + (Im z)\<^sup>2 = 0\ by (simp add: gauss_eq_iff sum_power2_eq_zero_iff) lemma gauss_neq_0: \z \ 0 \ (Re z)\<^sup>2 + (Im z)\<^sup>2 > 0\ by (simp add: gauss_eq_0 sum_power2_ge_zero less_le) 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 instance gauss :: idom proof fix x y :: gauss assume \x \ 0\ \y \ 0\ then show \x * y \ 0\ by (simp_all add: gauss_eq_iff) (smt (verit, best) mult_eq_0_iff mult_neg_neg mult_neg_pos mult_pos_neg mult_pos_pos) qed subsection \The Gauss Number $i$\ primcorec imaginary_unit :: gauss (\\\) where \Re \ = 0\ | \Im \ = 1\ lemma Gauss_eq: \Gauss a b = of_int a + \ * of_int b\ by (simp add: gauss_eq_iff) lemma gauss_eq: \a = of_int (Re a) + \ * of_int (Im a)\ by (simp add: gauss_eq_iff) lemma gauss_i_not_zero [simp]: \\ \ 0\ by (simp add: gauss_eq_iff) lemma gauss_i_not_one [simp]: \\ \ 1\ by (simp add: gauss_eq_iff) lemma gauss_i_not_numeral [simp]: \\ \ numeral n\ by (simp add: gauss_eq_iff) lemma gauss_i_not_neg_numeral [simp]: \\ \ - numeral n\ by (simp add: gauss_eq_iff) lemma i_mult_i_eq [simp]: \\ * \ = - 1\ by (simp add: gauss_eq_iff) lemma gauss_i_mult_minus [simp]: \\ * (\ * x) = - x\ by (simp flip: mult.assoc) lemma i_squared [simp]: \\\<^sup>2 = - 1\ by (simp add: power2_eq_square) lemma i_even_power [simp]: \\ ^ (n * 2) = (- 1) ^ n\ unfolding mult.commute [of n] power_mult by simp 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 is_unit_i [simp]: \\ dvd 1\ by (rule dvdI [of _ _ \- \\]) simp lemma gauss_numeral [code_post]: \Gauss 0 0 = 0\ \Gauss 1 0 = 1\ \Gauss (- 1) 0 = - 1\ \Gauss (numeral n) 0 = numeral n\ \Gauss (- numeral n) 0 = - numeral n\ \Gauss 0 1 = \\ \Gauss 0 (- 1) = - \\ \Gauss 0 (numeral n) = numeral n * \\ \Gauss 0 (- numeral n) = - numeral n * \\ \Gauss 1 1 = 1 + \\ \Gauss (- 1) 1 = - 1 + \\ \Gauss (numeral n) 1 = numeral n + \\ \Gauss (- numeral n) 1 = - numeral n + \\ \Gauss 1 (- 1) = 1 - \\ \Gauss 1 (numeral n) = 1 + numeral n * \\ \Gauss 1 (- numeral n) = 1 - numeral n * \\ \Gauss (- 1) (- 1) = - 1 - \\ \Gauss (numeral n) (- 1) = numeral n - \\ \Gauss (- numeral n) (- 1) = - numeral n - \\ \Gauss (- 1) (numeral n) = - 1 + numeral n * \\ \Gauss (- 1) (- numeral n) = - 1 - numeral n * \\ \Gauss (numeral m) (numeral n) = numeral m + numeral n * \\ \Gauss (- numeral m) (numeral n) = - numeral m + numeral n * \\ \Gauss (numeral m) (- numeral n) = numeral m - numeral n * \\ \Gauss (- numeral m) (- numeral n) = - numeral m - numeral n * \\ by (simp_all add: gauss_eq_iff) subsection \Gauss Conjugation\ primcorec cnj :: \gauss \ gauss\ where \Re (cnj z) = Re z\ | \Im (cnj z) = - Im z\ lemma gauss_cnj_cancel_iff [simp]: \cnj x = cnj y \ x = y\ by (simp add: gauss_eq_iff) lemma gauss_cnj_cnj [simp]: \cnj (cnj z) = z\ by (simp add: gauss_eq_iff) lemma gauss_cnj_zero [simp]: \cnj 0 = 0\ by (simp add: gauss_eq_iff) lemma gauss_cnj_zero_iff [iff]: \cnj z = 0 \ z = 0\ by (simp add: gauss_eq_iff) lemma gauss_cnj_one_iff [simp]: \cnj z = 1 \ z = 1\ by (simp add: gauss_eq_iff) lemma gauss_cnj_add [simp]: \cnj (x + y) = cnj x + cnj y\ by (simp add: gauss_eq_iff) lemma cnj_sum [simp]: \cnj (sum f s) = (\x\s. cnj (f x))\ by (induct s rule: infinite_finite_induct) auto lemma gauss_cnj_diff [simp]: \cnj (x - y) = cnj x - cnj y\ by (simp add: gauss_eq_iff) lemma gauss_cnj_minus [simp]: \cnj (- x) = - cnj x\ by (simp add: gauss_eq_iff) lemma gauss_cnj_one [simp]: \cnj 1 = 1\ by (simp add: gauss_eq_iff) lemma gauss_cnj_mult [simp]: \cnj (x * y) = cnj x * cnj y\ by (simp add: gauss_eq_iff) lemma cnj_prod [simp]: \cnj (prod f s) = (\x\s. cnj (f x))\ by (induct s rule: infinite_finite_induct) auto lemma gauss_cnj_power [simp]: \cnj (x ^ n) = cnj x ^ n\ by (induct n) simp_all lemma gauss_cnj_numeral [simp]: \cnj (numeral w) = numeral w\ by (simp add: gauss_eq_iff) lemma gauss_cnj_of_nat [simp]: \cnj (of_nat n) = of_nat n\ by (simp add: gauss_eq_iff) lemma gauss_cnj_of_int [simp]: \cnj (of_int z) = of_int z\ by (simp add: gauss_eq_iff) lemma gauss_cnj_i [simp]: \cnj \ = - \\ by (simp add: gauss_eq_iff) lemma gauss_add_cnj: \z + cnj z = of_int (2 * Re z)\ by (simp add: gauss_eq_iff) lemma gauss_diff_cnj: \z - cnj z = of_int (2 * Im z) * \\ by (simp add: gauss_eq_iff) lemma gauss_mult_cnj: \z * cnj z = of_int ((Re z)\<^sup>2 + (Im z)\<^sup>2)\ by (simp add: gauss_eq_iff power2_eq_square) lemma cnj_add_mult_eq_Re: \z * cnj w + cnj z * w = of_int (2 * Re (z * cnj w))\ by (simp add: gauss_eq_iff) lemma gauss_In_mult_cnj_zero [simp]: \Im (z * cnj z) = 0\ by simp subsection \Algebraic division\ instantiation gauss :: idom_modulo begin primcorec divide_gauss :: \gauss \ gauss \ gauss\ where - \Re (x div y) = (Re x * Re y + Im x * Im y) div ((Re y)\<^sup>2 + (Im y)\<^sup>2)\ - | \Im (x div y) = (Im x * Re y - Re x * Im y) div ((Re y)\<^sup>2 + (Im y)\<^sup>2)\ + \Re (x div y) = (Re x * Re y + Im x * Im y) rdiv ((Re y)\<^sup>2 + (Im y)\<^sup>2)\ + | \Im (x div y) = (Im x * Re y - Re x * Im y) rdiv ((Re y)\<^sup>2 + (Im y)\<^sup>2)\ definition modulo_gauss :: \gauss \ gauss \ gauss\ where \x mod y = x - x div y * y\ for x y :: gauss instance apply standard apply (simp_all add: modulo_gauss_def) apply (auto simp add: gauss_eq_iff algebra_simps power2_eq_square) apply (simp_all only: flip: mult.assoc distrib_right) apply (simp_all only: mult.assoc [of \Im k\ \Re l\ \Re r\ for k l r]) - apply (simp_all add: sum_squares_eq_zero_iff mult.commute flip: distrib_left) + apply (simp_all add: sum_squares_eq_zero_iff mult.commute nonzero_mult_rdiv_cancel_right flip: distrib_left) done end end diff --git a/src/HOL/Library/Library.thy b/src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy +++ b/src/HOL/Library/Library.thy @@ -1,101 +1,102 @@ (*<*) theory Library imports AList Adhoc_Overloading BigO BNF_Axiomatization BNF_Corec Bourbaki_Witt_Fixpoint Char_ord Code_Cardinality Code_Lazy Code_Test Combine_PER Complete_Partial_Order2 Conditional_Parametricity Confluence Confluent_Quotient Countable Countable_Complete_Lattices Countable_Set_Type Debug Diagonal_Subsequence Discrete Disjoint_Sets Disjoint_FSets Dlist Dual_Ordered_Lattice Equipollence Extended Extended_Nat Extended_Nonnegative_Real Extended_Real Finite_Map Float FSet FuncSet Function_Division Fun_Lexorder Going_To_Filter Groups_Big_Fun Indicator_Function Infinite_Set Interval Interval_Float IArray Landau_Symbols Lattice_Algebras Lattice_Constructions Linear_Temporal_Logic_on_Streams ListVector Lub_Glb Mapping Monad_Syntax More_List Multiset_Order NList Nonpos_Ints Numeral_Type Omega_Words_Fun Open_State_Syntax Option_ord Order_Continuity Parallel Pattern_Aliases Periodic_Fun Poly_Mapping Power_By_Squaring Preorder Product_Plus Quadratic_Discriminant Quotient_List Quotient_Option Quotient_Product Quotient_Set Quotient_Sum Quotient_Syntax Quotient_Type Ramsey Reflection Rewrite + Rounded_Division Saturated Set_Algebras Set_Idioms Signed_Division State_Monad Stream Sorting_Algorithms Sublist Sum_of_Squares Transitive_Closure_Table Tree_Multiset Tree_Real Type_Length Uprod While_Combinator Word Z2 begin end (*>*) diff --git a/src/HOL/Library/Rounded_Division.thy b/src/HOL/Library/Rounded_Division.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Library/Rounded_Division.thy @@ -0,0 +1,84 @@ +(* Author: Florian Haftmann, TU Muenchen +*) + +subsection \Rounded division: modulus centered towars zero.\ + +theory Rounded_Division + imports Main +begin + +definition rounded_divide :: \int \ int \ int\ (infixl \rdiv\ 70) + where \k rdiv l = (k + l div 2 + of_bool (l < 0)) div l\ + +definition rounded_modulo :: \int \ int \ int\ (infixl \rmod\ 70) + where \k rmod l = k - k rdiv l * l\ + +lemma rdiv_mult_rmod_eq: + \k rdiv l * l + k rmod l = k\ + by (simp add: rounded_divide_def rounded_modulo_def) + +lemma mult_rdiv_rmod_eq: + \l * (k rdiv l) + k rmod l = k\ + using rdiv_mult_rmod_eq [of k l] by (simp add: ac_simps) + +lemma rmod_rdiv_mult_eq: + \k rmod l + k rdiv l * l = k\ + using rdiv_mult_rmod_eq [of k l] by (simp add: ac_simps) + +lemma rmod_mult_rdiv_eq: + \k rmod l + l * (k rdiv l) = k\ + using rdiv_mult_rmod_eq [of k l] by (simp add: ac_simps) + +lemma minus_rdiv_mult_eq_rmod: + \k - k rdiv l * l = k rmod l\ + by (rule add_implies_diff [symmetric]) (fact rmod_rdiv_mult_eq) + +lemma minus_mult_rdiv_eq_rmod: + \k - l * (k rdiv l) = k rmod l\ + by (rule add_implies_diff [symmetric]) (fact rmod_mult_rdiv_eq) + +lemma minus_rmod_eq_rdiv_mult: + \k - k rmod l = k rdiv l * l\ + by (rule add_implies_diff [symmetric]) (fact rdiv_mult_rmod_eq) + +lemma minus_rmod_eq_mult_rdiv: + \k - k rmod l = l * (k rdiv l)\ + by (rule add_implies_diff [symmetric]) (fact mult_rdiv_rmod_eq) + +lemma rdiv_0_eq [simp]: + \k rdiv 0 = 0\ + by (simp add: rounded_divide_def) + +lemma rmod_0_eq [simp]: + \k rmod 0 = k\ + by (simp add: rounded_modulo_def) + +lemma rdiv_1_eq [simp]: + \k rdiv 1 = k\ + by (simp add: rounded_divide_def) + +lemma rmod_1_eq [simp]: + \k rmod 1 = 0\ + by (simp add: rounded_modulo_def) + +lemma zero_rdiv_eq [simp]: + \0 rdiv k = 0\ + by (auto simp add: rounded_divide_def not_less zdiv_eq_0_iff) + +lemma zero_rmod_eq [simp]: + \0 rmod k = 0\ + by (simp add: rounded_modulo_def) + +lemma nonzero_mult_rdiv_cancel_right: + \k * l rdiv l = k\ if \l \ 0\ + using that by (auto simp add: rounded_divide_def ac_simps) + +lemma rdiv_self_eq [simp]: + \k rdiv k = 1\ if \k \ 0\ + using that nonzero_mult_rdiv_cancel_right [of k 1] by simp + +lemma rmod_self_eq [simp]: + \k rmod k = 0\ + by (cases \k = 0\) (simp_all add: rounded_modulo_def) + +end