# HG changeset patch # User haftmann # Date 1718893726 0 # Node ID 31bf95336f16e73534ec5ef0082cb1ff959190b1 # Parent 898034c8a799efbea748c9d0911764dd9d1a6279 dropped references to theorems from transitional theory Divides.thy diff -r 898034c8a799 -r 31bf95336f16 src/Doc/Tutorial/Rules/Forward.thy --- a/src/Doc/Tutorial/Rules/Forward.thy Wed Jun 19 12:13:16 2024 +0200 +++ b/src/Doc/Tutorial/Rules/Forward.thy Thu Jun 20 14:28:46 2024 +0000 @@ -187,10 +187,6 @@ \rulename{div_mult_mod_eq} \ -(*MOVED to Force.thy, which now depends only on Divides.thy -lemma div_mult_self_is_m: "0 (m*n) div n = (m::nat)" -*) - lemma relprime_dvd_mult_iff: "gcd k n = 1 \ (k dvd m*n) = (k dvd m)" by (auto intro: relprime_dvd_mult elim: dvdE) diff -r 898034c8a799 -r 31bf95336f16 src/HOL/Library/Word.thy --- a/src/HOL/Library/Word.thy Wed Jun 19 12:13:16 2024 +0200 +++ b/src/HOL/Library/Word.thy Thu Jun 20 14:28:46 2024 +0000 @@ -1224,7 +1224,7 @@ end -context unique_euclidean_semiring_numeral +context linordered_euclidean_semiring begin lemma unsigned_greater_eq [simp]: @@ -2915,7 +2915,7 @@ by (metis mod_pos_pos_trivial uint_lt2p uint_mult_ge0 uint_word_ariths(3)) lemma uint_sub_lem: "uint x \ uint y \ uint (x - y) = uint x - uint y" - by (metis diff_ge_0_iff_ge of_nat_0_le_iff uint_nat uint_sub_lt2p uint_word_of_int unique_euclidean_semiring_numeral_class.mod_less word_sub_wi) + by (simp add: uint_word_arith_bintrs take_bit_int_eq_self_iff) lemma uint_add_le: "uint (x + y) \ uint x + uint y" unfolding uint_word_ariths by (simp add: zmod_le_nonneg_dividend) @@ -3728,6 +3728,15 @@ lemma word_rev_gal': "u = word_reverse w \ w = word_reverse u" by simp +lemma word_eq_reverseI: + \v = w\ if \word_reverse v = word_reverse w\ +proof - + from that have \word_reverse (word_reverse v) = word_reverse (word_reverse w)\ + by simp + then show ?thesis + by simp +qed + lemma uint_2p: "(0::'a::len word) < 2 ^ n \ uint (2 ^ n::'a::len word) = 2 ^ n" by (cases \n < LENGTH('a)\; transfer; force) @@ -4162,35 +4171,38 @@ \w = word_rotr n v \ v = word_rotl n w\ by auto -lemma word_rotr_rev: - \word_rotr n w = word_reverse (word_rotl n (word_reverse w))\ +lemma word_reverse_word_rotl: + \word_reverse (word_rotl n w) = word_rotr n (word_reverse w)\ (is \?lhs = ?rhs\) proof (rule bit_word_eqI) fix m assume \m < LENGTH('a)\ - moreover have \1 + - ((int m + int n mod int LENGTH('a)) mod int LENGTH('a) + - ((int LENGTH('a) * 2) mod int LENGTH('a) - (1 + (int m + int n mod int LENGTH('a)))) mod int LENGTH('a)) = - int LENGTH('a)\ - apply (cases \(1 + (int m + int n mod int LENGTH('a))) mod - int LENGTH('a) = 0\) - using zmod_zminus1_eq_if [of \1 + (int m + int n mod int LENGTH('a))\ \int LENGTH('a)\] - apply simp_all - apply (auto simp add: algebra_simps) - apply (metis (mono_tags, opaque_lifting) Abs_fnat_hom_add mod_Suc mod_mult_self2_is_0 of_nat_Suc of_nat_mod semiring_char_0_class.of_nat_neq_0) - apply (metis (no_types, opaque_lifting) Abs_fnat_hom_add less_not_refl mod_Suc of_nat_Suc of_nat_gt_0 of_nat_mod) + then have \int (LENGTH('a) - Suc ((m + n) mod LENGTH('a))) = + int ((LENGTH('a) + LENGTH('a) - Suc (m + n mod LENGTH('a))) mod LENGTH('a))\ + apply (simp only: of_nat_diff of_nat_mod) + apply (simp add: Suc_le_eq add_less_le_mono of_nat_mod algebra_simps) + apply (simp only: mod_diff_left_eq [symmetric, of \int LENGTH('a) * 2\] mod_mult_self1_is_0 diff_0 minus_mod_int_eq) + apply (simp add: mod_simps) done - then have \int ((m + n) mod LENGTH('a)) = - int (LENGTH('a) - Suc ((LENGTH('a) - Suc m + LENGTH('a) - n mod LENGTH('a)) mod LENGTH('a)))\ - using \m < LENGTH('a)\ - by (simp only: of_nat_mod mod_simps) - (simp add: of_nat_diff of_nat_mod Suc_le_eq add_less_mono algebra_simps mod_simps) - then have \(m + n) mod LENGTH('a) = - LENGTH('a) - Suc ((LENGTH('a) - Suc m + LENGTH('a) - n mod LENGTH('a)) mod LENGTH('a))\ + then have \LENGTH('a) - Suc ((m + n) mod LENGTH('a)) = + (LENGTH('a) + LENGTH('a) - Suc (m + n mod LENGTH('a))) mod + LENGTH('a)\ by simp - ultimately show \bit (word_rotr n w) m \ bit (word_reverse (word_rotl n (word_reverse w))) m\ - by (simp add: word_rotl_eq_word_rotr bit_word_rotr_iff bit_word_reverse_iff) + with \m < LENGTH('a)\ show \bit ?lhs m \ bit ?rhs m\ + by (simp add: bit_simps) qed +lemma word_reverse_word_rotr: + \word_reverse (word_rotr n w) = word_rotl n (word_reverse w)\ + by (rule word_eq_reverseI) (simp add: word_reverse_word_rotl) + +lemma word_rotl_rev: + \word_rotl n w = word_reverse (word_rotr n (word_reverse w))\ + by (simp add: word_reverse_word_rotr) + +lemma word_rotr_rev: + \word_rotr n w = word_reverse (word_rotl n (word_reverse w))\ + by (simp add: word_reverse_word_rotl) + lemma word_roti_0 [simp]: "word_roti 0 w = w" by transfer simp diff -r 898034c8a799 -r 31bf95336f16 src/HOL/Number_Theory/Gauss.thy --- a/src/HOL/Number_Theory/Gauss.thy Wed Jun 19 12:13:16 2024 +0200 +++ b/src/HOL/Number_Theory/Gauss.thy Thu Jun 20 14:28:46 2024 +0000 @@ -196,10 +196,7 @@ by (auto simp add: C_def B_mod_greater_zero) lemma F_subset: "F \ {x. 0 < x \ x \ ((int p - 1) div 2)}" - apply (auto simp add: F_def E_def C_def) - apply (metis p_ge_2 Divides.pos_mod_bound nat_int zless_nat_conj) - apply (auto intro: p_odd_int) - done + using p_ge_2 by (auto simp add: F_def E_def C_def intro: p_odd_int) lemma D_subset: "D \ {x. 0 < x \ x \ ((p - 1) div 2)}" by (auto simp add: D_def C_greater_zero) diff -r 898034c8a799 -r 31bf95336f16 src/HOL/Number_Theory/Quadratic_Reciprocity.thy --- a/src/HOL/Number_Theory/Quadratic_Reciprocity.thy Wed Jun 19 12:13:16 2024 +0200 +++ b/src/HOL/Number_Theory/Quadratic_Reciprocity.thy Thu Jun 20 14:28:46 2024 +0000 @@ -189,10 +189,8 @@ using assms by (simp add: mod_mod_cancel yk(1)) moreover have "(y mod (int p * int q)) mod int q = snd res" using assms by (simp add: mod_mod_cancel yk(2)) - ultimately obtain x where "P_1 res x" - unfolding P_1_def - using Divides.pos_mod_bound Divides.pos_mod_sign pq_ge_0 - by (metis atLeastAtMost_iff zle_diff1_eq) + ultimately have "P_1 res (int y mod (int p * int q))" + using pq_ge_0 by (simp add: P_1_def) moreover have "a = b" if "P_1 res a" "P_1 res b" for a b proof - from that have "int p * int q dvd a - b"