diff -r 72b78ee82f7b -r 930abfdf8727 src/HOL/Number_Theory/Cong.thy --- a/src/HOL/Number_Theory/Cong.thy Fri Oct 20 23:29:43 2017 +0200 +++ b/src/HOL/Number_Theory/Cong.thy Fri Oct 20 20:57:55 2017 +0200 @@ -32,236 +32,187 @@ imports "HOL-Computational_Algebra.Primes" begin -subsection \Turn off \One_nat_def\\ - -lemma power_eq_one_eq_nat [simp]: "x^m = 1 \ m = 0 \ x = 1" - for x m :: nat - by (induct m) auto - -declare mod_pos_pos_trivial [simp] - - -subsection \Main definitions\ - -class cong = - fixes cong :: "'a \ 'a \ 'a \ bool" ("(1[_ = _] '(()mod _'))") +subsection \Generic congruences\ + +context unique_euclidean_semiring begin +definition cong :: "'a \ 'a \ 'a \ bool" ("(1[_ = _] '(()mod _'))") + where "cong b c a \ b mod a = c mod a" + abbreviation notcong :: "'a \ 'a \ 'a \ bool" ("(1[_ \ _] '(()mod _'))") - where "notcong x y m \ \ cong x y m" + where "notcong b c a \ \ cong b c a" + +lemma cong_mod_left [simp]: + "[b mod a = c] (mod a) \ [b = c] (mod a)" + by (simp add: cong_def) + +lemma cong_mod_right [simp]: + "[b = c mod a] (mod a) \ [b = c] (mod a)" + by (simp add: cong_def) + +lemma cong_dvd_iff: + "a dvd b \ a dvd c" if "[b = c] (mod a)" + using that by (auto simp: cong_def dvd_eq_mod_eq_0) + +lemma cong_0 [simp, presburger]: + "[b = c] (mod 0) \ b = c" + by (simp add: cong_def) + +lemma cong_1 [simp, presburger]: + "[b = c] (mod 1)" + by (simp add: cong_def) + +lemma cong_refl [simp]: + "[b = b] (mod a)" + by (simp add: cong_def) + +lemma cong_sym: + "[b = c] (mod a) \ [c = b] (mod a)" + by (simp add: cong_def) + +lemma cong_sym_eq: + "[b = c] (mod a) \ [c = b] (mod a)" + by (auto simp add: cong_def) + +lemma cong_trans [trans]: + "[b = c] (mod a) \ [c = d] (mod a) \ [b = d] (mod a)" + by (simp add: cong_def) + +lemma cong_add: + "[b = c] (mod a) \ [d = e] (mod a) \ [b + d = c + e] (mod a)" + by (auto simp add: cong_def intro: mod_add_cong) + +lemma cong_mult: + "[b = c] (mod a) \ [d = e] (mod a) \ [b * d = c * e] (mod a)" + by (auto simp add: cong_def intro: mod_mult_cong) + +lemma cong_pow: + "[b = c] (mod a) \ [b ^ n = c ^ n] (mod a)" + by (simp add: cong_def power_mod [symmetric, of b n a] power_mod [symmetric, of c n a]) + +lemma cong_sum: + "[sum f A = sum g A] (mod a)" if "\x. x \ A \ [f x = g x] (mod a)" + using that by (induct A rule: infinite_finite_induct) (auto intro: cong_add) + +lemma cong_prod: + "[prod f A = prod g A] (mod a)" if "(\x. x \ A \ [f x = g x] (mod a))" + using that by (induct A rule: infinite_finite_induct) (auto intro: cong_mult) + +lemma cong_scalar_right: + "[b = c] (mod a) \ [b * d = c * d] (mod a)" + by (simp add: cong_mult) + +lemma cong_scalar_left: + "[b = c] (mod a) \ [d * b = d * c] (mod a)" + by (simp add: cong_mult) + +lemma cong_mult_self_right: "[b * a = 0] (mod a)" + by (simp add: cong_def) + +lemma cong_mult_self_left: "[a * b = 0] (mod a)" + by (simp add: cong_def) + +lemma cong_0_iff: "[b = 0] (mod a) \ a dvd b" + by (simp add: cong_def dvd_eq_mod_eq_0) + +lemma mod_mult_cong_right: + "[c mod (a * b) = d] (mod a) \ [c = d] (mod a)" + by (simp add: cong_def mod_mod_cancel mod_add_left_eq) + +lemma mod_mult_cong_left: + "[c mod (b * a) = d] (mod a) \ [c = d] (mod a)" + using mod_mult_cong_right [of c a b d] by (simp add: ac_simps) + +end + +context unique_euclidean_ring +begin + +lemma cong_diff: + "[b = c] (mod a) \ [d = e] (mod a) \ [b - d = c - e] (mod a)" + by (auto simp add: cong_def intro: mod_diff_cong) + +lemma cong_diff_iff_cong_0: + "[b - c = 0] (mod a) \ [b = c] (mod a)" (is "?P \ ?Q") +proof + assume ?P + then have "[b - c + c = 0 + c] (mod a)" + by (rule cong_add) simp + then show ?Q + by simp +next + assume ?Q + with cong_diff [of b c a c c] show ?P + by simp +qed + +lemma cong_minus_minus_iff: + "[- b = - c] (mod a) \ [b = c] (mod a)" + using cong_diff_iff_cong_0 [of b c a] cong_diff_iff_cong_0 [of "- b" "- c" a] + by (simp add: cong_0_iff dvd_diff_commute) + +lemma cong_modulus_minus_iff: "[b = c] (mod - a) \ [b = c] (mod a)" + using cong_diff_iff_cong_0 [of b c a] cong_diff_iff_cong_0 [of b c " -a"] + by (simp add: cong_0_iff) end -subsubsection \Definitions for the natural numbers\ - -instantiation nat :: cong -begin - -definition cong_nat :: "nat \ nat \ nat \ bool" - where "cong_nat x y m \ x mod m = y mod m" - -instance .. - -end - - -subsubsection \Definitions for the integers\ - -instantiation int :: cong -begin - -definition cong_int :: "int \ int \ int \ bool" - where "cong_int x y m \ x mod m = y mod m" - -instance .. - -end - - -subsection \Set up Transfer\ - +subsection \Congruences on @{typ nat} and @{typ int}\ lemma transfer_nat_int_cong: "x \ 0 \ y \ 0 \ m \ 0 \ [nat x = nat y] (mod (nat m)) \ [x = y] (mod m)" for x y m :: int - unfolding cong_int_def cong_nat_def - by (metis int_nat_eq nat_mod_distrib zmod_int) + by (auto simp add: cong_def nat_mod_distrib [symmetric]) + (metis eq_nat_nat_iff le_less mod_by_0 pos_mod_conj) declare transfer_morphism_nat_int [transfer add return: transfer_nat_int_cong] -lemma transfer_int_nat_cong: "[int x = int y] (mod (int m)) = [x = y] (mod m)" - by (auto simp add: cong_int_def cong_nat_def) (auto simp add: zmod_int [symmetric]) +lemma cong_int_iff: + "[int m = int q] (mod int n) \ [m = q] (mod n)" + by (simp add: cong_def of_nat_mod [symmetric]) + +lemma transfer_int_nat_cong: + "[int x = int y] (mod (int m)) = [x = y] (mod m)" + by (fact cong_int_iff) declare transfer_morphism_int_nat [transfer add return: transfer_int_nat_cong] - -subsection \Congruence\ - -(* was zcong_0, etc. *) -lemma cong_0_nat [simp, presburger]: "[a = b] (mod 0) \ a = b" - for a b :: nat - by (auto simp: cong_nat_def) - -lemma cong_0_int [simp, presburger]: "[a = b] (mod 0) \ a = b" - for a b :: int - by (auto simp: cong_int_def) - -lemma cong_1_nat [simp, presburger]: "[a = b] (mod 1)" - for a b :: nat - by (auto simp: cong_nat_def) - -lemma cong_Suc_0_nat [simp, presburger]: "[a = b] (mod Suc 0)" - for a b :: nat - by (auto simp: cong_nat_def) - -lemma cong_1_int [simp, presburger]: "[a = b] (mod 1)" - for a b :: int - by (auto simp: cong_int_def) - -lemma cong_refl_nat [simp]: "[k = k] (mod m)" - for k :: nat - by (auto simp: cong_nat_def) - -lemma cong_refl_int [simp]: "[k = k] (mod m)" - for k :: int - by (auto simp: cong_int_def) - -lemma cong_sym_nat: "[a = b] (mod m) \ [b = a] (mod m)" - for a b :: nat - by (auto simp: cong_nat_def) - -lemma cong_sym_int: "[a = b] (mod m) \ [b = a] (mod m)" - for a b :: int - by (auto simp: cong_int_def) - -lemma cong_sym_eq_nat: "[a = b] (mod m) = [b = a] (mod m)" - for a b :: nat - by (auto simp: cong_nat_def) - -lemma cong_sym_eq_int: "[a = b] (mod m) = [b = a] (mod m)" - for a b :: int - by (auto simp: cong_int_def) - -lemma cong_trans_nat [trans]: "[a = b] (mod m) \ [b = c] (mod m) \ [a = c] (mod m)" - for a b c :: nat - by (auto simp: cong_nat_def) - -lemma cong_trans_int [trans]: "[a = b] (mod m) \ [b = c] (mod m) \ [a = c] (mod m)" - for a b c :: int - by (auto simp: cong_int_def) - -lemma cong_add_nat: "[a = b] (mod m) \ [c = d] (mod m) \ [a + c = b + d] (mod m)" - for a b c :: nat - unfolding cong_nat_def by (metis mod_add_cong) - -lemma cong_add_int: "[a = b] (mod m) \ [c = d] (mod m) \ [a + c = b + d] (mod m)" - for a b c :: int - unfolding cong_int_def by (metis mod_add_cong) - -lemma cong_diff_int: "[a = b] (mod m) \ [c = d] (mod m) \ [a - c = b - d] (mod m)" - for a b c :: int - unfolding cong_int_def by (metis mod_diff_cong) +lemma cong_Suc_0 [simp, presburger]: + "[m = n] (mod Suc 0)" + using cong_1 [of m n] by simp lemma cong_diff_aux_int: "[a = b] (mod m) \ [c = d] (mod m) \ a \ c \ b \ d \ [tsub a c = tsub b d] (mod m)" for a b c d :: int - by (metis cong_diff_int tsub_eq) + by (metis cong_diff tsub_eq) lemma cong_diff_nat: - fixes a b c d :: nat - assumes "[a = b] (mod m)" "[c = d] (mod m)" "a \ c" "b \ d" - shows "[a - c = b - d] (mod m)" - using assms by (rule cong_diff_aux_int [transferred]) - -lemma cong_mult_nat: "[a = b] (mod m) \ [c = d] (mod m) \ [a * c = b * d] (mod m)" - for a b c d :: nat - unfolding cong_nat_def by (metis mod_mult_cong) - -lemma cong_mult_int: "[a = b] (mod m) \ [c = d] (mod m) \ [a * c = b * d] (mod m)" - for a b c d :: int - unfolding cong_int_def by (metis mod_mult_cong) - -lemma cong_exp_nat: "[x = y] (mod n) \ [x^k = y^k] (mod n)" - for x y :: nat - by (induct k) (auto simp: cong_mult_nat) - -lemma cong_exp_int: "[x = y] (mod n) \ [x^k = y^k] (mod n)" - for x y :: int - by (induct k) (auto simp: cong_mult_int) - -lemma cong_sum_nat: "(\x. x \ A \ [f x = g x] (mod m)) \ [(\x\A. f x) = (\x\A. g x)] (mod m)" - for f g :: "'a \ nat" - by (induct A rule: infinite_finite_induct) (auto intro: cong_add_nat) - -lemma cong_sum_int: "(\x. x \ A \ [f x = g x] (mod m)) \ [(\x\A. f x) = (\x\A. g x)] (mod m)" - for f g :: "'a \ int" - by (induct A rule: infinite_finite_induct) (auto intro: cong_add_int) - -lemma cong_prod_nat: "(\x. x \ A \ [f x = g x] (mod m)) \ [(\x\A. f x) = (\x\A. g x)] (mod m)" - for f g :: "'a \ nat" - by (induct A rule: infinite_finite_induct) (auto intro: cong_mult_nat) + "[a - c = b - d] (mod m)" if "[a = b] (mod m)" "[c = d] (mod m)" + and "a \ c" "b \ d" for a b c d m :: nat + using that by (rule cong_diff_aux_int [transferred]) -lemma cong_prod_int: "(\x. x \ A \ [f x = g x] (mod m)) \ [(\x\A. f x) = (\x\A. g x)] (mod m)" - for f g :: "'a \ int" - by (induct A rule: infinite_finite_induct) (auto intro: cong_mult_int) - -lemma cong_scalar_nat: "[a = b] (mod m) \ [a * k = b * k] (mod m)" - for a b k :: nat - by (rule cong_mult_nat) simp_all - -lemma cong_scalar_int: "[a = b] (mod m) \ [a * k = b * k] (mod m)" - for a b k :: int - by (rule cong_mult_int) simp_all - -lemma cong_scalar2_nat: "[a = b] (mod m) \ [k * a = k * b] (mod m)" - for a b k :: nat - by (rule cong_mult_nat) simp_all +lemma cong_diff_iff_cong_0_aux_int: "a \ b \ [tsub a b = 0] (mod m) = [a = b] (mod m)" + for a b :: int + by (subst tsub_eq, assumption, rule cong_diff_iff_cong_0) -lemma cong_scalar2_int: "[a = b] (mod m) \ [k * a = k * b] (mod m)" - for a b k :: int - by (rule cong_mult_int) simp_all - -lemma cong_mult_self_nat: "[a * m = 0] (mod m)" - for a m :: nat - by (auto simp: cong_nat_def) - -lemma cong_mult_self_int: "[a * m = 0] (mod m)" - for a m :: int - by (auto simp: cong_int_def) - -lemma cong_eq_diff_cong_0_int: "[a = b] (mod m) = [a - b = 0] (mod m)" - for a b :: int - by (metis cong_add_int cong_diff_int cong_refl_int diff_add_cancel diff_self) - -lemma cong_eq_diff_cong_0_aux_int: "a \ b \ [a = b] (mod m) = [tsub a b = 0] (mod m)" - for a b :: int - by (subst tsub_eq, assumption, rule cong_eq_diff_cong_0_int) - -lemma cong_eq_diff_cong_0_nat: +lemma cong_diff_iff_cong_0_nat: fixes a b :: nat assumes "a \ b" - shows "[a = b] (mod m) = [a - b = 0] (mod m)" - using assms by (rule cong_eq_diff_cong_0_aux_int [transferred]) - -lemma cong_diff_cong_0'_nat: - "[x = y] (mod n) \ (if x \ y then [y - x = 0] (mod n) else [x - y = 0] (mod n))" - for x y :: nat - by (metis cong_eq_diff_cong_0_nat cong_sym_nat nat_le_linear) + shows "[a - b = 0] (mod m) = [a = b] (mod m)" + using assms by (rule cong_diff_iff_cong_0_aux_int [transferred]) lemma cong_altdef_nat: "a \ b \ [a = b] (mod m) \ m dvd (a - b)" for a b :: nat - apply (subst cong_eq_diff_cong_0_nat, assumption) - apply (unfold cong_nat_def) - apply (simp add: dvd_eq_mod_eq_0 [symmetric]) - done + by (simp add: cong_0_iff [symmetric] cong_diff_iff_cong_0_nat) lemma cong_altdef_int: "[a = b] (mod m) \ m dvd (a - b)" for a b :: int - by (metis cong_int_def mod_eq_dvd_iff) + by (simp add: cong_0_iff [symmetric] cong_diff_iff_cong_0) -lemma cong_abs_int: "[x = y] (mod abs m) \ [x = y] (mod m)" +lemma cong_abs_int [simp]: "[x = y] (mod abs m) \ [x = y] (mod m)" for x y :: int by (simp add: cong_altdef_int) @@ -289,7 +240,6 @@ for a k m :: int by (simp add: mult.commute cong_mult_rcancel_int) -(* was zcong_zgcd_zmult_zmod *) lemma coprime_cong_mult_int: "[a = b] (mod m) \ [a = b] (mod n) \ coprime m n \ [a = b] (mod m * n)" for a b :: int @@ -302,19 +252,19 @@ lemma cong_less_imp_eq_nat: "0 \ a \ a < m \ 0 \ b \ b < m \ [a = b] (mod m) \ a = b" for a b :: nat - by (auto simp add: cong_nat_def) + by (auto simp add: cong_def) lemma cong_less_imp_eq_int: "0 \ a \ a < m \ 0 \ b \ b < m \ [a = b] (mod m) \ a = b" for a b :: int - by (auto simp add: cong_int_def) + by (auto simp add: cong_def) lemma cong_less_unique_nat: "0 < m \ (\!b. 0 \ b \ b < m \ [a = b] (mod m))" for a m :: nat - by (auto simp: cong_nat_def) (metis mod_less_divisor mod_mod_trivial) + by (auto simp: cong_def) (metis mod_less_divisor mod_mod_trivial) lemma cong_less_unique_int: "0 < m \ (\!b. 0 \ b \ b < m \ [a = b] (mod m))" for a m :: int - by (auto simp: cong_int_def) (metis mod_mod_trivial pos_mod_conj) + by (auto simp: cong_def) (metis mod_mod_trivial pos_mod_conj) lemma cong_iff_lin_int: "[a = b] (mod m) \ (\k. b = a + m * k)" for a b :: int @@ -334,20 +284,17 @@ next case False with \?lhs\ show ?rhs - apply (subst (asm) cong_sym_eq_nat) - apply (auto simp: cong_altdef_nat) - apply (metis add_0_right add_diff_inverse dvd_div_mult_self less_or_eq_imp_le mult_0) - done + by (metis cong_def mult.commute nat_le_linear nat_mod_eq_lemma) qed next assume ?rhs then show ?lhs - by (metis cong_nat_def mod_mult_self2 mult.commute) + by (metis cong_def mult.commute nat_mod_eq_iff) qed lemma cong_gcd_eq_int: "[a = b] (mod m) \ gcd a m = gcd b m" for a b :: int - by (metis cong_int_def gcd_red_int) + by (auto simp add: cong_def) (metis gcd_red_int) lemma cong_gcd_eq_nat: "[a = b] (mod m) \ gcd a m = gcd b m" for a b :: nat @@ -363,11 +310,11 @@ lemma cong_cong_mod_nat: "[a = b] (mod m) \ [a mod m = b mod m] (mod m)" for a b :: nat - by (auto simp add: cong_nat_def) + by simp lemma cong_cong_mod_int: "[a = b] (mod m) \ [a mod m = b mod m] (mod m)" for a b :: int - by (auto simp add: cong_int_def) + by simp lemma cong_minus_int [iff]: "[a = b] (mod - m) \ [a = b] (mod m)" for a b :: int @@ -434,53 +381,6 @@ for x y :: int by (auto simp add: cong_altdef_int dvd_def) -lemma cong_dvd_eq_nat: "[x = y] (mod n) \ n dvd x \ n dvd y" - for x y :: nat - by (auto simp: cong_nat_def dvd_eq_mod_eq_0) - -lemma cong_dvd_eq_int: "[x = y] (mod n) \ n dvd x \ n dvd y" - for x y :: int - by (auto simp: cong_int_def dvd_eq_mod_eq_0) - -lemma cong_mod_nat: "n \ 0 \ [a mod n = a] (mod n)" - for a n :: nat - by (simp add: cong_nat_def) - -lemma cong_mod_int: "n \ 0 \ [a mod n = a] (mod n)" - for a n :: int - by (simp add: cong_int_def) - -lemma mod_mult_cong_nat: "a \ 0 \ b \ 0 \ [x mod (a * b) = y] (mod a) \ [x = y] (mod a)" - for a b :: nat - by (simp add: cong_nat_def mod_mult2_eq mod_add_left_eq) - -lemma neg_cong_int: "[a = b] (mod m) \ [- a = - b] (mod m)" - for a b :: int - by (metis cong_int_def minus_minus mod_minus_cong) - -lemma cong_modulus_neg_int: "[a = b] (mod m) \ [a = b] (mod - m)" - for a b :: int - by (auto simp add: cong_altdef_int) - -lemma mod_mult_cong_int: "a \ 0 \ b \ 0 \ [x mod (a * b) = y] (mod a) \ [x = y] (mod a)" - for a b :: int -proof (cases "b > 0") - case True - then show ?thesis - by (simp add: cong_int_def mod_mod_cancel mod_add_left_eq) -next - case False - then show ?thesis - apply (subst (1 2) cong_modulus_neg_int) - apply (unfold cong_int_def) - apply (subgoal_tac "a * b = (- a * - b)") - apply (erule ssubst) - apply (subst zmod_zmult2_eq) - apply (auto simp add: mod_add_left_eq mod_minus_right div_minus_right) - apply (metis mod_diff_left_eq mod_diff_right_eq mod_mult_self1_is_0 diff_zero)+ - done -qed - lemma cong_to_1_nat: fixes a :: nat assumes "[a = 1] (mod n)" @@ -494,15 +394,15 @@ qed lemma cong_0_1_nat': "[0 = Suc 0] (mod n) \ n = Suc 0" - by (auto simp: cong_nat_def) + by (auto simp: cong_def) lemma cong_0_1_nat: "[0 = 1] (mod n) \ n = 1" for n :: nat - by (auto simp: cong_nat_def) + by (auto simp: cong_def) lemma cong_0_1_int: "[0 = 1] (mod n) \ n = 1 \ n = - 1" for n :: int - by (auto simp: cong_int_def zmult_eq_1_iff) + by (auto simp: cong_def zmult_eq_1_iff) lemma cong_to_1'_nat: "[a = 1] (mod n) \ a = 0 \ n = 1 \ (\m. a = 1 + m * n)" for a :: nat @@ -524,7 +424,7 @@ case False then show ?thesis using bezout_nat [of a n, OF \a \ 0\] - by auto (metis cong_add_rcancel_0_nat cong_mult_self_nat mult.commute) + by auto (metis cong_add_rcancel_0_nat cong_mult_self_left) qed lemma cong_solve_int: "a \ 0 \ \x. [a * x = gcd a n] (mod n)" @@ -546,7 +446,7 @@ from cong_solve_nat [OF a] obtain x where "[a * x = gcd a n](mod n)" by auto then have "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)" - by (elim cong_scalar2_nat) + using cong_scalar_left by blast also from b have "(d div gcd a n) * gcd a n = d" by (rule dvd_div_mult_self) also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)" @@ -562,7 +462,7 @@ from cong_solve_int [OF a] obtain x where "[a * x = gcd a n](mod n)" by auto then have "[(d div gcd a n) * (a * x) = (d div gcd a n) * gcd a n] (mod n)" - by (elim cong_scalar2_int) + using cong_scalar_left by blast also from b have "(d div gcd a n) * gcd a n = d" by (rule dvd_div_mult_self) also have "(d div gcd a n) * (a * x) = a * (d div gcd a n * x)" @@ -577,10 +477,12 @@ shows "\x. [a * x = 1] (mod n)" proof (cases "a = 0") case True - with assms show ?thesis by force + with assms show ?thesis + by (simp add: cong_0_1_nat') next case False - with assms show ?thesis by (metis cong_solve_nat) + with assms show ?thesis + by (metis cong_solve_nat) qed lemma cong_solve_coprime_int: "coprime (a::int) n \ \x. [a * x = 1] (mod n)" @@ -598,14 +500,14 @@ lemma coprime_iff_invertible_int: "m > 0 \ coprime a m \ (\x. [a * x = 1] (mod m))" for m :: int apply (auto intro: cong_solve_coprime_int) - apply (metis cong_int_def coprime_mul_eq gcd_1_int gcd.commute gcd_red_int) + using cong_gcd_eq_int coprime_mul_eq' apply fastforce done lemma coprime_iff_invertible'_nat: "m > 0 \ coprime a m \ (\x. 0 \ x \ x < m \ [a * x = Suc 0] (mod m))" apply (subst coprime_iff_invertible_nat) apply auto - apply (auto simp add: cong_nat_def) + apply (auto simp add: cong_def) apply (metis mod_less_divisor mod_mult_right_eq) done @@ -613,35 +515,35 @@ "m > 0 \ coprime a m \ (\x. 0 \ x \ x < m \ [a * x = 1] (mod m))" for m :: int apply (subst coprime_iff_invertible_int) - apply (auto simp add: cong_int_def) + apply (auto simp add: cong_def) apply (metis mod_mult_right_eq pos_mod_conj) done lemma cong_cong_lcm_nat: "[x = y] (mod a) \ [x = y] (mod b) \ [x = y] (mod lcm a b)" for x y :: nat apply (cases "y \ x") - apply (metis cong_altdef_nat lcm_least) - apply (meson cong_altdef_nat cong_sym_nat lcm_least_iff nat_le_linear) + apply (simp add: cong_altdef_nat) + apply (meson cong_altdef_nat cong_sym lcm_least_iff nat_le_linear) done lemma cong_cong_lcm_int: "[x = y] (mod a) \ [x = y] (mod b) \ [x = y] (mod lcm a b)" for x y :: int by (auto simp add: cong_altdef_int lcm_least) -lemma cong_cong_prod_coprime_nat [rule_format]: "finite A \ - (\i\A. (\j\A. i \ j \ coprime (m i) (m j))) \ - (\i\A. [(x::nat) = y] (mod m i)) \ - [x = y] (mod (\i\A. m i))" - apply (induct set: finite) - apply auto +lemma cong_cong_prod_coprime_nat: + "[x = y] (mod (\i\A. m i))" if + "(\i\A. [(x::nat) = y] (mod m i))" + and "(\i\A. (\j\A. i \ j \ coprime (m i) (m j)))" + using that apply (induct A rule: infinite_finite_induct) + apply auto apply (metis One_nat_def coprime_cong_mult_nat gcd.commute prod_coprime) done -lemma cong_cong_prod_coprime_int [rule_format]: "finite A \ - (\i\A. (\j\A. i \ j \ coprime (m i) (m j))) \ - (\i\A. [(x::int) = y] (mod m i)) \ - [x = y] (mod (\i\A. m i))" - apply (induct set: finite) +lemma cong_cong_prod_coprime_int [rule_format]: + "[x = y] (mod (\i\A. m i))" if + "(\i\A. [(x::int) = y] (mod m i))" + "(\i\A. (\j\A. i \ j \ coprime (m i) (m j)))" + using that apply (induct A rule: infinite_finite_induct) apply auto apply (metis coprime_cong_mult_int gcd.commute prod_coprime) done @@ -658,9 +560,9 @@ from cong_solve_coprime_nat [OF b] obtain x2 where 2: "[m2 * x2 = 1] (mod m1)" by auto have "[m1 * x1 = 0] (mod m1)" - by (subst mult.commute) (rule cong_mult_self_nat) + by (simp add: cong_mult_self_left) moreover have "[m2 * x2 = 0] (mod m2)" - by (subst mult.commute) (rule cong_mult_self_nat) + by (simp add: cong_mult_self_left) ultimately show ?thesis using 1 2 by blast qed @@ -677,9 +579,9 @@ from cong_solve_coprime_int [OF b] obtain x2 where 2: "[m2 * x2 = 1] (mod m1)" by auto have "[m1 * x1 = 0] (mod m1)" - by (subst mult.commute) (rule cong_mult_self_int) + by (simp add: cong_mult_self_left) moreover have "[m2 * x2 = 0] (mod m2)" - by (subst mult.commute) (rule cong_mult_self_int) + by (simp add: cong_mult_self_left) ultimately show ?thesis using 1 2 by blast qed @@ -695,20 +597,10 @@ by blast let ?x = "u1 * b1 + u2 * b2" have "[?x = u1 * 1 + u2 * 0] (mod m1)" - apply (rule cong_add_nat) - apply (rule cong_scalar2_nat) - apply (rule \[b1 = 1] (mod m1)\) - apply (rule cong_scalar2_nat) - apply (rule \[b2 = 0] (mod m1)\) - done + using \[b1 = 1] (mod m1)\ \[b2 = 0] (mod m1)\ cong_add cong_scalar_left by blast then have "[?x = u1] (mod m1)" by simp have "[?x = u1 * 0 + u2 * 1] (mod m2)" - apply (rule cong_add_nat) - apply (rule cong_scalar2_nat) - apply (rule \[b1 = 0] (mod m2)\) - apply (rule cong_scalar2_nat) - apply (rule \[b2 = 1] (mod m2)\) - done + using \[b1 = 0] (mod m2)\ \[b2 = 1] (mod m2)\ cong_add cong_scalar_left by blast then have "[?x = u2] (mod m2)" by simp with \[?x = u1] (mod m1)\ show ?thesis @@ -726,20 +618,10 @@ by blast let ?x = "u1 * b1 + u2 * b2" have "[?x = u1 * 1 + u2 * 0] (mod m1)" - apply (rule cong_add_int) - apply (rule cong_scalar2_int) - apply (rule \[b1 = 1] (mod m1)\) - apply (rule cong_scalar2_int) - apply (rule \[b2 = 0] (mod m1)\) - done + using \[b1 = 1] (mod m1)\ \[b2 = 0] (mod m1)\ cong_add cong_scalar_left by blast then have "[?x = u1] (mod m1)" by simp have "[?x = u1 * 0 + u2 * 1] (mod m2)" - apply (rule cong_add_int) - apply (rule cong_scalar2_int) - apply (rule \[b1 = 0] (mod m2)\) - apply (rule cong_scalar2_int) - apply (rule \[b2 = 1] (mod m2)\) - done + using \[b1 = 0] (mod m2)\ \[b2 = 1] (mod m2)\ cong_add cong_scalar_left by blast then have "[?x = u2] (mod m2)" by simp with \[?x = u1] (mod m1)\ show ?thesis by blast @@ -750,8 +632,8 @@ apply (cases "y \ x") apply (simp add: cong_altdef_nat) apply (erule dvd_mult_left) - apply (rule cong_sym_nat) - apply (subst (asm) cong_sym_eq_nat) + apply (rule cong_sym) + apply (subst (asm) cong_sym_eq) apply (simp add: cong_altdef_nat) apply (erule dvd_mult_left) done @@ -764,7 +646,7 @@ lemma cong_less_modulus_unique_nat: "[x = y] (mod m) \ x < m \ y < m \ x = y" for x y :: nat - by (simp add: cong_nat_def) + by (simp add: cong_def) lemma binary_chinese_remainder_unique_nat: fixes m1 m2 :: nat @@ -779,21 +661,19 @@ from nz have less: "?x < m1 * m2" by auto have 1: "[?x = u1] (mod m1)" - apply (rule cong_trans_nat) + apply (rule cong_trans) prefer 2 apply (rule \[y = u1] (mod m1)\) - apply (rule cong_modulus_mult_nat) - apply (rule cong_mod_nat) - using nz apply auto + apply (rule cong_modulus_mult_nat [of _ _ _ m2]) + apply simp done have 2: "[?x = u2] (mod m2)" - apply (rule cong_trans_nat) + apply (rule cong_trans) prefer 2 apply (rule \[y = u2] (mod m2)\) apply (subst mult.commute) - apply (rule cong_modulus_mult_nat) - apply (rule cong_mod_nat) - using nz apply auto + apply (rule cong_modulus_mult_nat [of _ _ _ m1]) + apply simp done have "\z. z < m1 * m2 \ [z = u1] (mod m1) \ [z = u2] (mod m2) \ z = ?x" proof clarify @@ -801,22 +681,22 @@ assume "z < m1 * m2" assume "[z = u1] (mod m1)" and "[z = u2] (mod m2)" have "[?x = z] (mod m1)" - apply (rule cong_trans_nat) + apply (rule cong_trans) apply (rule \[?x = u1] (mod m1)\) - apply (rule cong_sym_nat) + apply (rule cong_sym) apply (rule \[z = u1] (mod m1)\) done moreover have "[?x = z] (mod m2)" - apply (rule cong_trans_nat) + apply (rule cong_trans) apply (rule \[?x = u2] (mod m2)\) - apply (rule cong_sym_nat) + apply (rule cong_sym) apply (rule \[z = u2] (mod m2)\) done ultimately have "[?x = z] (mod m1 * m2)" - by (auto intro: coprime_cong_mult_nat a) + using a by (auto intro: coprime_cong_mult_nat simp add: mod_mult_cong_left mod_mult_cong_right) with \z < m1 * m2\ \?x < m1 * m2\ show "z = ?x" apply (intro cong_less_modulus_unique_nat) - apply (auto, erule cong_sym_nat) + apply (auto, erule cong_sym) done qed with less 1 2 show ?thesis by auto @@ -838,7 +718,7 @@ then obtain x where "[(\j \ A - {i}. m j) * x = 1] (mod m i)" by auto moreover have "[(\j \ A - {i}. m j) * x = 0] (mod (\j \ A - {i}. m j))" - by (subst mult.commute, rule cong_mult_self_nat) + by (simp add: cong_0_iff) ultimately show "\a. [a = 1] (mod m i) \ [a = 0] (mod prod m (A - {i}))" by blast qed @@ -867,11 +747,11 @@ by auto also have "[u i * b i + (\j \ A - {i}. u j * b j) = u i * 1 + (\j \ A - {i}. u j * 0)] (mod m i)" - apply (rule cong_add_nat) - apply (rule cong_scalar2_nat) + apply (rule cong_add) + apply (rule cong_scalar_left) using b a apply blast - apply (rule cong_sum_nat) - apply (rule cong_scalar2_nat) + apply (rule cong_sum) + apply (rule cong_scalar_left) using b apply auto apply (rule cong_dvd_modulus_nat) apply (drule (1) bspec) @@ -886,11 +766,11 @@ qed qed -lemma coprime_cong_prod_nat [rule_format]: "finite A \ - (\i\A. (\j\A. i \ j \ coprime (m i) (m j))) \ - (\i\A. [(x::nat) = y] (mod m i)) \ - [x = y] (mod (\i\A. m i))" - apply (induct set: finite) +lemma coprime_cong_prod_nat: + "[x = y] (mod (\i\A. m i))" + if "\i\A. (\j\A. i \ j \ coprime (m i) (m j))" + and "\i\A. [x = y] (mod m i)" for x y :: nat + using that apply (induct A rule: infinite_finite_induct) apply auto apply (metis One_nat_def coprime_cong_mult_nat gcd.commute prod_coprime) done @@ -914,15 +794,12 @@ by auto have cong: "\i\A. [?x = u i] (mod m i)" apply auto - apply (rule cong_trans_nat) + apply (rule cong_trans) prefer 2 using one apply auto - apply (rule cong_dvd_modulus_nat) - apply (rule cong_mod_nat) - using prodnz apply auto - apply rule - apply (rule fin) - apply assumption + apply (rule cong_dvd_modulus_nat [of _ _ "prod m A"]) + apply simp + using fin apply auto done have unique: "\z. z < (\i\A. m i) \ (\i\A. [z = u i] (mod m i)) \ z = ?x" proof clarify @@ -931,9 +808,9 @@ assume zcong: "(\i\A. [z = u i] (mod m i))" have "\i\A. [?x = z] (mod m i)" apply clarify - apply (rule cong_trans_nat) + apply (rule cong_trans) using cong apply (erule bspec) - apply (rule cong_sym_nat) + apply (rule cong_sym) using zcong apply auto done with fin cop have "[?x = z] (mod (\i\A. m i))" @@ -943,7 +820,7 @@ with zless less show "z = ?x" apply (intro cong_less_modulus_unique_nat) apply auto - apply (erule cong_sym_nat) + apply (erule cong_sym) done qed from less cong unique show ?thesis