diff --git a/src/HOL/Parity.thy b/src/HOL/Parity.thy --- a/src/HOL/Parity.thy +++ b/src/HOL/Parity.thy @@ -1,1100 +1,1104 @@ (* Title: HOL/Parity.thy Author: Jeremy Avigad Author: Jacques D. Fleuriot *) section \Parity in rings and semirings\ theory Parity imports Euclidean_Division begin subsection \Ring structures with parity and \even\/\odd\ predicates\ class semiring_parity = comm_semiring_1 + semiring_modulo + assumes even_iff_mod_2_eq_zero: "2 dvd a \ a mod 2 = 0" and odd_iff_mod_2_eq_one: "\ 2 dvd a \ a mod 2 = 1" and odd_one [simp]: "\ 2 dvd 1" begin abbreviation even :: "'a \ bool" where "even a \ 2 dvd a" abbreviation odd :: "'a \ bool" where "odd a \ \ 2 dvd a" lemma parity_cases [case_names even odd]: assumes "even a \ a mod 2 = 0 \ P" assumes "odd a \ a mod 2 = 1 \ P" shows P using assms by (cases "even a") (simp_all add: even_iff_mod_2_eq_zero [symmetric] odd_iff_mod_2_eq_one [symmetric]) lemma not_mod_2_eq_0_eq_1 [simp]: "a mod 2 \ 0 \ a mod 2 = 1" by (cases a rule: parity_cases) simp_all lemma not_mod_2_eq_1_eq_0 [simp]: "a mod 2 \ 1 \ a mod 2 = 0" by (cases a rule: parity_cases) simp_all lemma mod2_eq_if: "a mod 2 = (if 2 dvd a then 0 else 1)" by (simp add: even_iff_mod_2_eq_zero odd_iff_mod_2_eq_one) lemma evenE [elim?]: assumes "even a" obtains b where "a = 2 * b" using assms by (rule dvdE) lemma oddE [elim?]: assumes "odd a" obtains b where "a = 2 * b + 1" proof - have "a = 2 * (a div 2) + a mod 2" by (simp add: mult_div_mod_eq) with assms have "a = 2 * (a div 2) + 1" by (simp add: odd_iff_mod_2_eq_one) then show ?thesis .. qed lemma mod_2_eq_odd: "a mod 2 = of_bool (odd a)" by (auto elim: oddE simp add: even_iff_mod_2_eq_zero) lemma of_bool_odd_eq_mod_2: "of_bool (odd a) = a mod 2" by (simp add: mod_2_eq_odd) lemma even_zero [simp]: "even 0" by (fact dvd_0_right) lemma odd_even_add: "even (a + b)" if "odd a" and "odd b" proof - from that obtain c d where "a = 2 * c + 1" and "b = 2 * d + 1" by (blast elim: oddE) then have "a + b = 2 * c + 2 * d + (1 + 1)" by (simp only: ac_simps) also have "\ = 2 * (c + d + 1)" by (simp add: algebra_simps) finally show ?thesis .. qed lemma even_add [simp]: "even (a + b) \ (even a \ even b)" by (auto simp add: dvd_add_right_iff dvd_add_left_iff odd_even_add) lemma odd_add [simp]: "odd (a + b) \ \ (odd a \ odd b)" by simp lemma even_plus_one_iff [simp]: "even (a + 1) \ odd a" by (auto simp add: dvd_add_right_iff intro: odd_even_add) lemma even_mult_iff [simp]: "even (a * b) \ even a \ even b" (is "?P \ ?Q") proof assume ?Q then show ?P by auto next assume ?P show ?Q proof (rule ccontr) assume "\ (even a \ even b)" then have "odd a" and "odd b" by auto then obtain r s where "a = 2 * r + 1" and "b = 2 * s + 1" by (blast elim: oddE) then have "a * b = (2 * r + 1) * (2 * s + 1)" by simp also have "\ = 2 * (2 * r * s + r + s) + 1" by (simp add: algebra_simps) finally have "odd (a * b)" by simp with \?P\ show False by auto qed qed lemma even_numeral [simp]: "even (numeral (Num.Bit0 n))" proof - have "even (2 * numeral n)" unfolding even_mult_iff by simp then have "even (numeral n + numeral n)" unfolding mult_2 . then show ?thesis unfolding numeral.simps . qed lemma odd_numeral [simp]: "odd (numeral (Num.Bit1 n))" proof assume "even (numeral (num.Bit1 n))" then have "even (numeral n + numeral n + 1)" unfolding numeral.simps . then have "even (2 * numeral n + 1)" unfolding mult_2 . then have "2 dvd numeral n * 2 + 1" by (simp add: ac_simps) then have "2 dvd 1" using dvd_add_times_triv_left_iff [of 2 "numeral n" 1] by simp then show False by simp qed lemma even_power [simp]: "even (a ^ n) \ even a \ n > 0" by (induct n) auto end class ring_parity = ring + semiring_parity begin subclass comm_ring_1 .. lemma even_minus: "even (- a) \ even a" by (fact dvd_minus_iff) lemma even_diff [simp]: "even (a - b) \ even (a + b)" using even_add [of a "- b"] by simp end subsection \Special case: euclidean rings containing the natural numbers\ class unique_euclidean_semiring_with_nat = semidom + semiring_char_0 + unique_euclidean_semiring + assumes of_nat_div: "of_nat (m div n) = of_nat m div of_nat n" and division_segment_of_nat [simp]: "division_segment (of_nat n) = 1" and division_segment_euclidean_size [simp]: "division_segment a * of_nat (euclidean_size a) = a" begin lemma division_segment_eq_iff: "a = b" if "division_segment a = division_segment b" and "euclidean_size a = euclidean_size b" using that division_segment_euclidean_size [of a] by simp lemma euclidean_size_of_nat [simp]: "euclidean_size (of_nat n) = n" proof - have "division_segment (of_nat n) * of_nat (euclidean_size (of_nat n)) = of_nat n" by (fact division_segment_euclidean_size) then show ?thesis by simp qed lemma of_nat_euclidean_size: "of_nat (euclidean_size a) = a div division_segment a" proof - have "of_nat (euclidean_size a) = division_segment a * of_nat (euclidean_size a) div division_segment a" by (subst nonzero_mult_div_cancel_left) simp_all also have "\ = a div division_segment a" by simp finally show ?thesis . qed lemma division_segment_1 [simp]: "division_segment 1 = 1" using division_segment_of_nat [of 1] by simp lemma division_segment_numeral [simp]: "division_segment (numeral k) = 1" using division_segment_of_nat [of "numeral k"] by simp lemma euclidean_size_1 [simp]: "euclidean_size 1 = 1" using euclidean_size_of_nat [of 1] by simp lemma euclidean_size_numeral [simp]: "euclidean_size (numeral k) = numeral k" using euclidean_size_of_nat [of "numeral k"] by simp lemma of_nat_dvd_iff: "of_nat m dvd of_nat n \ m dvd n" (is "?P \ ?Q") proof (cases "m = 0") case True then show ?thesis by simp next case False show ?thesis proof assume ?Q then show ?P by auto next assume ?P with False have "of_nat n = of_nat n div of_nat m * of_nat m" by simp then have "of_nat n = of_nat (n div m * m)" by (simp add: of_nat_div) then have "n = n div m * m" by (simp only: of_nat_eq_iff) then have "n = m * (n div m)" by (simp add: ac_simps) then show ?Q .. qed qed lemma of_nat_mod: "of_nat (m mod n) = of_nat m mod of_nat n" proof - have "of_nat m div of_nat n * of_nat n + of_nat m mod of_nat n = of_nat m" by (simp add: div_mult_mod_eq) also have "of_nat m = of_nat (m div n * n + m mod n)" by simp finally show ?thesis by (simp only: of_nat_div of_nat_mult of_nat_add) simp qed lemma one_div_two_eq_zero [simp]: "1 div 2 = 0" proof - from of_nat_div [symmetric] have "of_nat 1 div of_nat 2 = of_nat 0" by (simp only:) simp then show ?thesis by simp qed lemma one_mod_two_eq_one [simp]: "1 mod 2 = 1" proof - from of_nat_mod [symmetric] have "of_nat 1 mod of_nat 2 = of_nat 1" by (simp only:) simp then show ?thesis by simp qed subclass semiring_parity proof show "2 dvd a \ a mod 2 = 0" for a by (fact dvd_eq_mod_eq_0) show "\ 2 dvd a \ a mod 2 = 1" for a proof assume "a mod 2 = 1" then show "\ 2 dvd a" by auto next assume "\ 2 dvd a" have eucl: "euclidean_size (a mod 2) = 1" proof (rule order_antisym) show "euclidean_size (a mod 2) \ 1" using mod_size_less [of 2 a] by simp show "1 \ euclidean_size (a mod 2)" using \\ 2 dvd a\ by (simp add: Suc_le_eq dvd_eq_mod_eq_0) qed from \\ 2 dvd a\ have "\ of_nat 2 dvd division_segment a * of_nat (euclidean_size a)" by simp then have "\ of_nat 2 dvd of_nat (euclidean_size a)" by (auto simp only: dvd_mult_unit_iff' is_unit_division_segment) then have "\ 2 dvd euclidean_size a" using of_nat_dvd_iff [of 2] by simp then have "euclidean_size a mod 2 = 1" by (simp add: semidom_modulo_class.dvd_eq_mod_eq_0) then have "of_nat (euclidean_size a mod 2) = of_nat 1" by simp then have "of_nat (euclidean_size a) mod 2 = 1" by (simp add: of_nat_mod) from \\ 2 dvd a\ eucl show "a mod 2 = 1" by (auto intro: division_segment_eq_iff simp add: division_segment_mod) qed show "\ is_unit 2" proof (rule notI) assume "is_unit 2" then have "of_nat 2 dvd of_nat 1" by simp then have "is_unit (2::nat)" by (simp only: of_nat_dvd_iff) then show False by simp qed qed lemma even_of_nat [simp]: "even (of_nat a) \ even a" proof - have "even (of_nat a) \ of_nat 2 dvd of_nat a" by simp also have "\ \ even a" by (simp only: of_nat_dvd_iff) finally show ?thesis . qed lemma one_mod_2_pow_eq [simp]: "1 mod (2 ^ n) = of_bool (n > 0)" proof - have "1 mod (2 ^ n) = of_nat (1 mod (2 ^ n))" using of_nat_mod [of 1 "2 ^ n"] by simp also have "\ = of_bool (n > 0)" by simp finally show ?thesis . qed lemma one_div_2_pow_eq [simp]: "1 div (2 ^ n) = of_bool (n = 0)" using div_mult_mod_eq [of 1 "2 ^ n"] by auto lemma even_succ_div_two [simp]: "even a \ (a + 1) div 2 = a div 2" by (cases "a = 0") (auto elim!: evenE dest: mult_not_zero) lemma odd_succ_div_two [simp]: "odd a \ (a + 1) div 2 = a div 2 + 1" by (auto elim!: oddE simp add: add.assoc) lemma even_two_times_div_two: "even a \ 2 * (a div 2) = a" by (fact dvd_mult_div_cancel) lemma odd_two_times_div_two_succ [simp]: "odd a \ 2 * (a div 2) + 1 = a" using mult_div_mod_eq [of 2 a] by (simp add: even_iff_mod_2_eq_zero) lemma coprime_left_2_iff_odd [simp]: "coprime 2 a \ odd a" proof assume "odd a" show "coprime 2 a" proof (rule coprimeI) fix b assume "b dvd 2" "b dvd a" then have "b dvd a mod 2" by (auto intro: dvd_mod) with \odd a\ show "is_unit b" by (simp add: mod_2_eq_odd) qed next assume "coprime 2 a" show "odd a" proof (rule notI) assume "even a" then obtain b where "a = 2 * b" .. with \coprime 2 a\ have "coprime 2 (2 * b)" by simp moreover have "\ coprime 2 (2 * b)" by (rule not_coprimeI [of 2]) simp_all ultimately show False by blast qed qed lemma coprime_right_2_iff_odd [simp]: "coprime a 2 \ odd a" using coprime_left_2_iff_odd [of a] by (simp add: ac_simps) lemma div_mult2_eq': "a div (of_nat m * of_nat n) = a div of_nat m div of_nat n" proof (cases a "of_nat m * of_nat n" rule: divmod_cases) case (divides q) then show ?thesis using nonzero_mult_div_cancel_right [of "of_nat m" "q * of_nat n"] by (simp add: ac_simps) next case (remainder q r) then have "division_segment r = 1" using division_segment_of_nat [of "m * n"] by simp with division_segment_euclidean_size [of r] have "of_nat (euclidean_size r) = r" by simp have "a mod (of_nat m * of_nat n) div (of_nat m * of_nat n) = 0" by simp with remainder(6) have "r div (of_nat m * of_nat n) = 0" by simp with \of_nat (euclidean_size r) = r\ have "of_nat (euclidean_size r) div (of_nat m * of_nat n) = 0" by simp then have "of_nat (euclidean_size r div (m * n)) = 0" by (simp add: of_nat_div) then have "of_nat (euclidean_size r div m div n) = 0" by (simp add: div_mult2_eq) with \of_nat (euclidean_size r) = r\ have "r div of_nat m div of_nat n = 0" by (simp add: of_nat_div) with remainder(1) have "q = (r div of_nat m + q * of_nat n * of_nat m div of_nat m) div of_nat n" by simp with remainder(5) remainder(7) show ?thesis using div_plus_div_distrib_dvd_right [of "of_nat m" "q * (of_nat m * of_nat n)" r] by (simp add: ac_simps) next case by0 then show ?thesis by auto qed lemma mod_mult2_eq': "a mod (of_nat m * of_nat n) = of_nat m * (a div of_nat m mod of_nat n) + a mod of_nat m" proof - have "a div (of_nat m * of_nat n) * (of_nat m * of_nat n) + a mod (of_nat m * of_nat n) = a div of_nat m div of_nat n * of_nat n * of_nat m + (a div of_nat m mod of_nat n * of_nat m + a mod of_nat m)" by (simp add: combine_common_factor div_mult_mod_eq) moreover have "a div of_nat m div of_nat n * of_nat n * of_nat m = of_nat n * of_nat m * (a div of_nat m div of_nat n)" by (simp add: ac_simps) ultimately show ?thesis by (simp add: div_mult2_eq' mult_commute) qed lemma div_mult2_numeral_eq: "a div numeral k div numeral l = a div numeral (k * l)" (is "?A = ?B") proof - have "?A = a div of_nat (numeral k) div of_nat (numeral l)" by simp also have "\ = a div (of_nat (numeral k) * of_nat (numeral l))" by (fact div_mult2_eq' [symmetric]) also have "\ = ?B" by simp finally show ?thesis . qed lemma numeral_Bit0_div_2: "numeral (num.Bit0 n) div 2 = numeral n" proof - have "numeral (num.Bit0 n) = numeral n + numeral n" by (simp only: numeral.simps) also have "\ = numeral n * 2" by (simp add: mult_2_right) finally have "numeral (num.Bit0 n) div 2 = numeral n * 2 div 2" by simp also have "\ = numeral n" by (rule nonzero_mult_div_cancel_right) simp finally show ?thesis . qed lemma numeral_Bit1_div_2: "numeral (num.Bit1 n) div 2 = numeral n" proof - have "numeral (num.Bit1 n) = numeral n + numeral n + 1" by (simp only: numeral.simps) also have "\ = numeral n * 2 + 1" by (simp add: mult_2_right) finally have "numeral (num.Bit1 n) div 2 = (numeral n * 2 + 1) div 2" by simp also have "\ = numeral n * 2 div 2 + 1 div 2" using dvd_triv_right by (rule div_plus_div_distrib_dvd_left) also have "\ = numeral n * 2 div 2" by simp also have "\ = numeral n" by (rule nonzero_mult_div_cancel_right) simp finally show ?thesis . qed end class unique_euclidean_ring_with_nat = ring + unique_euclidean_semiring_with_nat begin subclass ring_parity .. lemma minus_1_mod_2_eq [simp]: "- 1 mod 2 = 1" by (simp add: mod_2_eq_odd) lemma minus_1_div_2_eq [simp]: "- 1 div 2 = - 1" proof - from div_mult_mod_eq [of "- 1" 2] have "- 1 div 2 * 2 = - 1 * 2" using add_implies_diff by fastforce then show ?thesis using mult_right_cancel [of 2 "- 1 div 2" "- 1"] by simp qed end subsection \Instance for \<^typ>\nat\\ instance nat :: unique_euclidean_semiring_with_nat by standard (simp_all add: dvd_eq_mod_eq_0) lemma even_Suc_Suc_iff [simp]: "even (Suc (Suc n)) \ even n" using dvd_add_triv_right_iff [of 2 n] by simp lemma even_Suc [simp]: "even (Suc n) \ odd n" using even_plus_one_iff [of n] by simp lemma even_diff_nat [simp]: "even (m - n) \ m < n \ even (m + n)" for m n :: nat proof (cases "n \ m") case True then have "m - n + n * 2 = m + n" by (simp add: mult_2_right) moreover have "even (m - n) \ even (m - n + n * 2)" by simp ultimately have "even (m - n) \ even (m + n)" by (simp only:) then show ?thesis by auto next case False then show ?thesis by simp qed lemma odd_pos: "odd n \ 0 < n" for n :: nat by (auto elim: oddE) lemma Suc_double_not_eq_double: "Suc (2 * m) \ 2 * n" proof assume "Suc (2 * m) = 2 * n" moreover have "odd (Suc (2 * m))" and "even (2 * n)" by simp_all ultimately show False by simp qed lemma double_not_eq_Suc_double: "2 * m \ Suc (2 * n)" using Suc_double_not_eq_double [of n m] by simp lemma odd_Suc_minus_one [simp]: "odd n \ Suc (n - Suc 0) = n" by (auto elim: oddE) lemma even_Suc_div_two [simp]: "even n \ Suc n div 2 = n div 2" using even_succ_div_two [of n] by simp lemma odd_Suc_div_two [simp]: "odd n \ Suc n div 2 = Suc (n div 2)" using odd_succ_div_two [of n] by simp lemma odd_two_times_div_two_nat [simp]: assumes "odd n" shows "2 * (n div 2) = n - (1 :: nat)" proof - from assms have "2 * (n div 2) + 1 = n" by (rule odd_two_times_div_two_succ) then have "Suc (2 * (n div 2)) - 1 = n - 1" by simp then show ?thesis by simp qed lemma not_mod2_eq_Suc_0_eq_0 [simp]: "n mod 2 \ Suc 0 \ n mod 2 = 0" using not_mod_2_eq_1_eq_0 [of n] by simp lemma nat_bit_induct [case_names zero even odd]: "P n" if zero: "P 0" and even: "\n. P n \ n > 0 \ P (2 * n)" and odd: "\n. P n \ P (Suc (2 * n))" proof (induction n rule: less_induct) case (less n) show "P n" proof (cases "n = 0") case True with zero show ?thesis by simp next case False with less have hyp: "P (n div 2)" by simp show ?thesis proof (cases "even n") case True then have "n \ 1" by auto with \n \ 0\ have "n div 2 > 0" by simp with \even n\ hyp even [of "n div 2"] show ?thesis by simp next case False with hyp odd [of "n div 2"] show ?thesis by simp qed qed qed lemma odd_card_imp_not_empty: \A \ {}\ if \odd (card A)\ using that by auto lemma nat_induct2 [case_names 0 1 step]: assumes "P 0" "P 1" and step: "\n::nat. P n \ P (n + 2)" shows "P n" proof (induct n rule: less_induct) case (less n) show ?case proof (cases "n < Suc (Suc 0)") case True then show ?thesis using assms by (auto simp: less_Suc_eq) next case False then obtain k where k: "n = Suc (Suc k)" by (force simp: not_less nat_le_iff_add) then have "kParity and powers\ context ring_1 begin lemma power_minus_even [simp]: "even n \ (- a) ^ n = a ^ n" by (auto elim: evenE) lemma power_minus_odd [simp]: "odd n \ (- a) ^ n = - (a ^ n)" by (auto elim: oddE) lemma uminus_power_if: "(- a) ^ n = (if even n then a ^ n else - (a ^ n))" by auto lemma neg_one_even_power [simp]: "even n \ (- 1) ^ n = 1" by simp lemma neg_one_odd_power [simp]: "odd n \ (- 1) ^ n = - 1" by simp lemma neg_one_power_add_eq_neg_one_power_diff: "k \ n \ (- 1) ^ (n + k) = (- 1) ^ (n - k)" by (cases "even (n + k)") auto lemma minus_one_power_iff: "(- 1) ^ n = (if even n then 1 else - 1)" by (induct n) auto end context linordered_idom begin lemma zero_le_even_power: "even n \ 0 \ a ^ n" by (auto elim: evenE) lemma zero_le_odd_power: "odd n \ 0 \ a ^ n \ 0 \ a" by (auto simp add: power_even_eq zero_le_mult_iff elim: oddE) lemma zero_le_power_eq: "0 \ a ^ n \ even n \ odd n \ 0 \ a" by (auto simp add: zero_le_even_power zero_le_odd_power) lemma zero_less_power_eq: "0 < a ^ n \ n = 0 \ even n \ a \ 0 \ odd n \ 0 < a" proof - have [simp]: "0 = a ^ n \ a = 0 \ n > 0" unfolding power_eq_0_iff [of a n, symmetric] by blast show ?thesis unfolding less_le zero_le_power_eq by auto qed lemma power_less_zero_eq [simp]: "a ^ n < 0 \ odd n \ a < 0" unfolding not_le [symmetric] zero_le_power_eq by auto lemma power_le_zero_eq: "a ^ n \ 0 \ n > 0 \ (odd n \ a \ 0 \ even n \ a = 0)" unfolding not_less [symmetric] zero_less_power_eq by auto lemma power_even_abs: "even n \ \a\ ^ n = a ^ n" using power_abs [of a n] by (simp add: zero_le_even_power) lemma power_mono_even: assumes "even n" and "\a\ \ \b\" shows "a ^ n \ b ^ n" proof - have "0 \ \a\" by auto with \\a\ \ \b\\ have "\a\ ^ n \ \b\ ^ n" by (rule power_mono) with \even n\ show ?thesis by (simp add: power_even_abs) qed lemma power_mono_odd: assumes "odd n" and "a \ b" shows "a ^ n \ b ^ n" proof (cases "b < 0") case True with \a \ b\ have "- b \ - a" and "0 \ - b" by auto then have "(- b) ^ n \ (- a) ^ n" by (rule power_mono) with \odd n\ show ?thesis by simp next case False then have "0 \ b" by auto show ?thesis proof (cases "a < 0") case True then have "n \ 0" and "a \ 0" using \odd n\ [THEN odd_pos] by auto then have "a ^ n \ 0" unfolding power_le_zero_eq using \odd n\ by auto moreover from \0 \ b\ have "0 \ b ^ n" by auto ultimately show ?thesis by auto next case False then have "0 \ a" by auto with \a \ b\ show ?thesis using power_mono by auto qed qed text \Simplify, when the exponent is a numeral\ lemma zero_le_power_eq_numeral [simp]: "0 \ a ^ numeral w \ even (numeral w :: nat) \ odd (numeral w :: nat) \ 0 \ a" by (fact zero_le_power_eq) lemma zero_less_power_eq_numeral [simp]: "0 < a ^ numeral w \ numeral w = (0 :: nat) \ even (numeral w :: nat) \ a \ 0 \ odd (numeral w :: nat) \ 0 < a" by (fact zero_less_power_eq) lemma power_le_zero_eq_numeral [simp]: "a ^ numeral w \ 0 \ (0 :: nat) < numeral w \ (odd (numeral w :: nat) \ a \ 0 \ even (numeral w :: nat) \ a = 0)" by (fact power_le_zero_eq) lemma power_less_zero_eq_numeral [simp]: "a ^ numeral w < 0 \ odd (numeral w :: nat) \ a < 0" by (fact power_less_zero_eq) lemma power_even_abs_numeral [simp]: "even (numeral w :: nat) \ \a\ ^ numeral w = a ^ numeral w" by (fact power_even_abs) end subsection \Instance for \<^typ>\int\\ instance int :: unique_euclidean_ring_with_nat by standard (simp_all add: dvd_eq_mod_eq_0 divide_int_def division_segment_int_def) lemma even_diff_iff: "even (k - l) \ even (k + l)" for k l :: int by (fact even_diff) lemma even_abs_add_iff: "even (\k\ + l) \ even (k + l)" for k l :: int by simp lemma even_add_abs_iff: "even (k + \l\) \ even (k + l)" for k l :: int by simp lemma even_nat_iff: "0 \ k \ even (nat k) \ even k" by (simp add: even_of_nat [of "nat k", where ?'a = int, symmetric]) lemma int_bit_induct [case_names zero minus even odd]: "P k" if zero_int: "P 0" and minus_int: "P (- 1)" and even_int: "\k. P k \ k \ 0 \ P (k * 2)" and odd_int: "\k. P k \ k \ - 1 \ P (1 + (k * 2))" for k :: int proof (cases "k \ 0") case True define n where "n = nat k" with True have "k = int n" by simp then show "P k" proof (induction n arbitrary: k rule: nat_bit_induct) case zero then show ?case by (simp add: zero_int) next case (even n) have "P (int n * 2)" by (rule even_int) (use even in simp_all) with even show ?case by (simp add: ac_simps) next case (odd n) have "P (1 + (int n * 2))" by (rule odd_int) (use odd in simp_all) with odd show ?case by (simp add: ac_simps) qed next case False define n where "n = nat (- k - 1)" with False have "k = - int n - 1" by simp then show "P k" proof (induction n arbitrary: k rule: nat_bit_induct) case zero then show ?case by (simp add: minus_int) next case (even n) have "P (1 + (- int (Suc n) * 2))" by (rule odd_int) (use even in \simp_all add: algebra_simps\) also have "\ = - int (2 * n) - 1" by (simp add: algebra_simps) finally show ?case using even by simp next case (odd n) have "P (- int (Suc n) * 2)" by (rule even_int) (use odd in \simp_all add: algebra_simps\) also have "\ = - int (Suc (2 * n)) - 1" by (simp add: algebra_simps) finally show ?case using odd by simp qed qed subsection \Abstract bit operations\ context semiring_parity begin text \The primary purpose of the following operations is to avoid ad-hoc simplification of concrete expressions \<^term>\2 ^ n\\ definition push_bit :: "nat \ 'a \ 'a" where push_bit_eq_mult: "push_bit n a = a * 2 ^ n" definition take_bit :: "nat \ 'a \ 'a" where take_bit_eq_mod: "take_bit n a = a mod 2 ^ n" definition drop_bit :: "nat \ 'a \ 'a" where drop_bit_eq_div: "drop_bit n a = a div 2 ^ n" end context unique_euclidean_semiring_with_nat begin lemma bit_ident: "push_bit n (drop_bit n a) + take_bit n a = a" using div_mult_mod_eq by (simp add: push_bit_eq_mult take_bit_eq_mod drop_bit_eq_div) lemma push_bit_push_bit [simp]: "push_bit m (push_bit n a) = push_bit (m + n) a" by (simp add: push_bit_eq_mult power_add ac_simps) lemma take_bit_take_bit [simp]: "take_bit m (take_bit n a) = take_bit (min m n) a" proof (cases "m \ n") case True then show ?thesis by (simp add: take_bit_eq_mod not_le min_def mod_mod_cancel le_imp_power_dvd) next case False then have "n < m" and "min m n = n" by simp_all then have "2 ^ m = of_nat (2 ^ n) * of_nat (2 ^ (m - n))" by (simp add: power_add [symmetric]) then have "a mod 2 ^ n mod 2 ^ m = a mod 2 ^ n mod (of_nat (2 ^ n) * of_nat (2 ^ (m - n)))" by simp also have "\ = of_nat (2 ^ n) * (a mod 2 ^ n div of_nat (2 ^ n) mod of_nat (2 ^ (m - n))) + a mod 2 ^ n mod of_nat (2 ^ n)" by (simp only: mod_mult2_eq') finally show ?thesis using \min m n = n\ by (simp add: take_bit_eq_mod) qed lemma drop_bit_drop_bit [simp]: "drop_bit m (drop_bit n a) = drop_bit (m + n) a" proof - have "a div (2 ^ m * 2 ^ n) = a div (of_nat (2 ^ n) * of_nat (2 ^ m))" by (simp add: ac_simps) also have "\ = a div of_nat (2 ^ n) div of_nat (2 ^ m)" by (simp only: div_mult2_eq') finally show ?thesis by (simp add: drop_bit_eq_div power_add) qed lemma push_bit_take_bit: "push_bit m (take_bit n a) = take_bit (m + n) (push_bit m a)" by (simp add: push_bit_eq_mult take_bit_eq_mod power_add mult_mod_right ac_simps) lemma take_bit_push_bit: "take_bit m (push_bit n a) = push_bit n (take_bit (m - n) a)" proof (cases "m \ n") case True then show ?thesis by (simp_all add: push_bit_eq_mult take_bit_eq_mod mod_eq_0_iff_dvd dvd_power_le) next case False then show ?thesis using push_bit_take_bit [of n "m - n" a] by simp qed lemma take_bit_drop_bit: "take_bit m (drop_bit n a) = drop_bit n (take_bit (m + n) a)" using mod_mult2_eq' [of a "2 ^ n" "2 ^ m"] by (simp add: drop_bit_eq_div take_bit_eq_mod power_add ac_simps) lemma drop_bit_take_bit: "drop_bit m (take_bit n a) = take_bit (n - m) (drop_bit m a)" proof (cases "m \ n") case True then show ?thesis using take_bit_drop_bit [of "n - m" m a] by simp next case False then have "a mod 2 ^ n div 2 ^ m = a mod 2 ^ n div 2 ^ (n + (m - n))" by simp also have "\ = a mod 2 ^ n div (2 ^ n * 2 ^ (m - n))" by (simp add: power_add) also have "\ = a mod 2 ^ n div (of_nat (2 ^ n) * of_nat (2 ^ (m - n)))" by simp also have "\ = a mod 2 ^ n div of_nat (2 ^ n) div of_nat (2 ^ (m - n))" by (simp only: div_mult2_eq') finally show ?thesis using False by (simp add: take_bit_eq_mod drop_bit_eq_div) qed lemma push_bit_0_id [simp]: "push_bit 0 = id" by (simp add: fun_eq_iff push_bit_eq_mult) lemma push_bit_of_0 [simp]: "push_bit n 0 = 0" by (simp add: push_bit_eq_mult) lemma push_bit_of_1: "push_bit n 1 = 2 ^ n" by (simp add: push_bit_eq_mult) lemma push_bit_Suc [simp]: "push_bit (Suc n) a = push_bit n (a * 2)" by (simp add: push_bit_eq_mult ac_simps) lemma push_bit_double: "push_bit n (a * 2) = push_bit n a * 2" by (simp add: push_bit_eq_mult ac_simps) lemma push_bit_eq_0_iff [simp]: "push_bit n a = 0 \ a = 0" by (simp add: push_bit_eq_mult) lemma push_bit_add: "push_bit n (a + b) = push_bit n a + push_bit n b" by (simp add: push_bit_eq_mult algebra_simps) lemma push_bit_numeral [simp]: "push_bit (numeral l) (numeral k) = push_bit (pred_numeral l) (numeral (Num.Bit0 k))" by (simp only: numeral_eq_Suc power_Suc numeral_Bit0 [of k] mult_2 [symmetric]) (simp add: ac_simps) lemma push_bit_of_nat: "push_bit n (of_nat m) = of_nat (push_bit n m)" by (simp add: push_bit_eq_mult Parity.push_bit_eq_mult) lemma take_bit_0 [simp]: "take_bit 0 a = 0" by (simp add: take_bit_eq_mod) lemma take_bit_Suc [simp]: "take_bit (Suc n) a = take_bit n (a div 2) * 2 + of_bool (odd a)" proof - have "1 + 2 * (a div 2) mod (2 * 2 ^ n) = (a div 2 * 2 + a mod 2) mod (2 * 2 ^ n)" if "odd a" using that mod_mult2_eq' [of "1 + 2 * (a div 2)" 2 "2 ^ n"] by (simp add: ac_simps odd_iff_mod_2_eq_one mult_mod_right) also have "\ = a mod (2 * 2 ^ n)" by (simp only: div_mult_mod_eq) finally show ?thesis by (simp add: take_bit_eq_mod algebra_simps mult_mod_right) qed lemma take_bit_of_0 [simp]: "take_bit n 0 = 0" by (simp add: take_bit_eq_mod) lemma take_bit_of_1 [simp]: "take_bit n 1 = of_bool (n > 0)" by (simp add: take_bit_eq_mod) lemma take_bit_add: "take_bit n (take_bit n a + take_bit n b) = take_bit n (a + b)" by (simp add: take_bit_eq_mod mod_simps) lemma take_bit_eq_0_iff: "take_bit n a = 0 \ 2 ^ n dvd a" by (simp add: take_bit_eq_mod mod_eq_0_iff_dvd) lemma take_bit_of_1_eq_0_iff [simp]: "take_bit n 1 = 0 \ n = 0" by (simp add: take_bit_eq_mod) lemma even_take_bit_eq [simp]: "even (take_bit n a) \ n = 0 \ even a" by (cases n) (simp_all add: take_bit_eq_mod dvd_mod_iff) lemma take_bit_numeral_bit0 [simp]: "take_bit (numeral l) (numeral (Num.Bit0 k)) = take_bit (pred_numeral l) (numeral k) * 2" by (simp only: numeral_eq_Suc power_Suc numeral_Bit0 [of k] mult_2 [symmetric] take_bit_Suc ac_simps even_mult_iff nonzero_mult_div_cancel_right [OF numeral_neq_zero]) simp lemma take_bit_numeral_bit1 [simp]: "take_bit (numeral l) (numeral (Num.Bit1 k)) = take_bit (pred_numeral l) (numeral k) * 2 + 1" by (simp only: numeral_eq_Suc power_Suc numeral_Bit1 [of k] mult_2 [symmetric] take_bit_Suc ac_simps even_add even_mult_iff div_mult_self1 [OF numeral_neq_zero]) (simp add: ac_simps) lemma take_bit_of_nat: "take_bit n (of_nat m) = of_nat (take_bit n m)" by (simp add: take_bit_eq_mod Parity.take_bit_eq_mod of_nat_mod [of m "2 ^ n"]) lemma drop_bit_0 [simp]: "drop_bit 0 = id" by (simp add: fun_eq_iff drop_bit_eq_div) lemma drop_bit_of_0 [simp]: "drop_bit n 0 = 0" by (simp add: drop_bit_eq_div) lemma drop_bit_of_1 [simp]: "drop_bit n 1 = of_bool (n = 0)" by (simp add: drop_bit_eq_div) lemma drop_bit_Suc [simp]: "drop_bit (Suc n) a = drop_bit n (a div 2)" proof (cases "even a") case True then obtain b where "a = 2 * b" .. moreover have "drop_bit (Suc n) (2 * b) = drop_bit n b" by (simp add: drop_bit_eq_div) ultimately show ?thesis by simp next case False then obtain b where "a = 2 * b + 1" .. moreover have "drop_bit (Suc n) (2 * b + 1) = drop_bit n b" using div_mult2_eq' [of "1 + b * 2" 2 "2 ^ n"] by (auto simp add: drop_bit_eq_div ac_simps) ultimately show ?thesis by simp qed lemma drop_bit_half: "drop_bit n (a div 2) = drop_bit n a div 2" by (induction n arbitrary: a) simp_all lemma drop_bit_of_bool [simp]: "drop_bit n (of_bool d) = of_bool (n = 0 \ d)" by (cases n) simp_all lemma drop_bit_numeral_bit0 [simp]: "drop_bit (numeral l) (numeral (Num.Bit0 k)) = drop_bit (pred_numeral l) (numeral k)" by (simp only: numeral_eq_Suc power_Suc numeral_Bit0 [of k] mult_2 [symmetric] drop_bit_Suc nonzero_mult_div_cancel_left [OF numeral_neq_zero]) lemma drop_bit_numeral_bit1 [simp]: "drop_bit (numeral l) (numeral (Num.Bit1 k)) = drop_bit (pred_numeral l) (numeral k)" by (simp only: numeral_eq_Suc power_Suc numeral_Bit1 [of k] mult_2 [symmetric] drop_bit_Suc div_mult_self4 [OF numeral_neq_zero]) simp lemma drop_bit_of_nat: "drop_bit n (of_nat m) = of_nat (drop_bit n m)" by (simp add: drop_bit_eq_div Parity.drop_bit_eq_div of_nat_div [of m "2 ^ n"]) end lemma push_bit_of_Suc_0 [simp]: "push_bit n (Suc 0) = 2 ^ n" using push_bit_of_1 [where ?'a = nat] by simp lemma take_bit_of_Suc_0 [simp]: "take_bit n (Suc 0) = of_bool (0 < n)" using take_bit_of_1 [where ?'a = nat] by simp lemma drop_bit_of_Suc_0 [simp]: "drop_bit n (Suc 0) = of_bool (n = 0)" using drop_bit_of_1 [where ?'a = nat] by simp +lemma take_bit_eq_self: + \take_bit n m = m\ if \m < 2 ^ n\ for n m :: nat + using that by (simp add: take_bit_eq_mod) + lemma push_bit_minus_one: "push_bit n (- 1 :: int) = - (2 ^ n)" by (simp add: push_bit_eq_mult) end diff --git a/src/HOL/ex/Word_Type.thy b/src/HOL/ex/Word_Type.thy --- a/src/HOL/ex/Word_Type.thy +++ b/src/HOL/ex/Word_Type.thy @@ -1,606 +1,730 @@ (* Author: Florian Haftmann, TUM *) section \Proof of concept for algebraically founded bit word types\ theory Word_Type imports Main "HOL-ex.Bit_Lists" "HOL-Library.Type_Length" begin subsection \Preliminaries\ lemma take_bit_uminus: "take_bit n (- (take_bit n k)) = take_bit n (- k)" for k :: int by (simp add: take_bit_eq_mod mod_minus_eq) lemma take_bit_minus: "take_bit n (take_bit n k - take_bit n l) = take_bit n (k - l)" for k l :: int by (simp add: take_bit_eq_mod mod_diff_eq) lemma take_bit_nonnegative [simp]: "take_bit n k \ 0" for k :: int by (simp add: take_bit_eq_mod) definition signed_take_bit :: "nat \ int \ int" where signed_take_bit_eq_take_bit: "signed_take_bit n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n" lemma signed_take_bit_eq_take_bit': "signed_take_bit (n - Suc 0) k = take_bit n (k + 2 ^ (n - 1)) - 2 ^ (n - 1)" if "n > 0" using that by (simp add: signed_take_bit_eq_take_bit) lemma signed_take_bit_0 [simp]: "signed_take_bit 0 k = - (k mod 2)" proof (cases "even k") case True then have "odd (k + 1)" by simp then have "(k + 1) mod 2 = 1" by (simp add: even_iff_mod_2_eq_zero) with True show ?thesis by (simp add: signed_take_bit_eq_take_bit) next case False then show ?thesis by (simp add: signed_take_bit_eq_take_bit odd_iff_mod_2_eq_one) qed lemma signed_take_bit_Suc [simp]: "signed_take_bit (Suc n) k = signed_take_bit n (k div 2) * 2 + k mod 2" by (simp add: odd_iff_mod_2_eq_one signed_take_bit_eq_take_bit algebra_simps) lemma signed_take_bit_of_0 [simp]: "signed_take_bit n 0 = 0" by (simp add: signed_take_bit_eq_take_bit take_bit_eq_mod) lemma signed_take_bit_of_minus_1 [simp]: "signed_take_bit n (- 1) = - 1" by (induct n) simp_all lemma signed_take_bit_eq_iff_take_bit_eq: "signed_take_bit (n - Suc 0) k = signed_take_bit (n - Suc 0) l \ take_bit n k = take_bit n l" (is "?P \ ?Q") if "n > 0" proof - from that obtain m where m: "n = Suc m" by (cases n) auto show ?thesis proof assume ?Q have "take_bit (Suc m) (k + 2 ^ m) = take_bit (Suc m) (take_bit (Suc m) k + take_bit (Suc m) (2 ^ m))" by (simp only: take_bit_add) also have "\ = take_bit (Suc m) (take_bit (Suc m) l + take_bit (Suc m) (2 ^ m))" by (simp only: \?Q\ m [symmetric]) also have "\ = take_bit (Suc m) (l + 2 ^ m)" by (simp only: take_bit_add) finally show ?P by (simp only: signed_take_bit_eq_take_bit m) simp next assume ?P with that have "(k + 2 ^ (n - Suc 0)) mod 2 ^ n = (l + 2 ^ (n - Suc 0)) mod 2 ^ n" by (simp add: signed_take_bit_eq_take_bit' take_bit_eq_mod) then have "(i + (k + 2 ^ (n - Suc 0))) mod 2 ^ n = (i + (l + 2 ^ (n - Suc 0))) mod 2 ^ n" for i by (metis mod_add_eq) then have "k mod 2 ^ n = l mod 2 ^ n" by (metis add_diff_cancel_right' uminus_add_conv_diff) then show ?Q by (simp add: take_bit_eq_mod) qed qed subsection \Bit strings as quotient type\ subsubsection \Basic properties\ quotient_type (overloaded) 'a word = int / "\k l. take_bit LENGTH('a) k = take_bit LENGTH('a::len0) l" by (auto intro!: equivpI reflpI sympI transpI) instantiation word :: (len0) "{semiring_numeral, comm_semiring_0, comm_ring}" begin lift_definition zero_word :: "'a word" is 0 . lift_definition one_word :: "'a word" is 1 . lift_definition plus_word :: "'a word \ 'a word \ 'a word" is plus by (subst take_bit_add [symmetric]) (simp add: take_bit_add) lift_definition uminus_word :: "'a word \ 'a word" is uminus by (subst take_bit_uminus [symmetric]) (simp add: take_bit_uminus) lift_definition minus_word :: "'a word \ 'a word \ 'a word" is minus by (subst take_bit_minus [symmetric]) (simp add: take_bit_minus) lift_definition times_word :: "'a word \ 'a word \ 'a word" is times by (auto simp add: take_bit_eq_mod intro: mod_mult_cong) instance by standard (transfer; simp add: algebra_simps)+ end instance word :: (len) comm_ring_1 by standard (transfer; simp)+ quickcheck_generator word constructors: "zero_class.zero :: ('a::len0) word", "numeral :: num \ ('a::len0) word", "uminus :: ('a::len0) word \ ('a::len0) word" +context + includes lifting_syntax + notes power_transfer [transfer_rule] +begin + +lemma power_transfer_word [transfer_rule]: + \(pcr_word ===> (=) ===> pcr_word) (^) (^)\ + by transfer_prover + +end + subsubsection \Conversions\ context includes lifting_syntax notes transfer_rule_numeral [transfer_rule] transfer_rule_of_nat [transfer_rule] transfer_rule_of_int [transfer_rule] begin lemma [transfer_rule]: "((=) ===> (pcr_word :: int \ 'a::len word \ bool)) numeral numeral" by transfer_prover lemma [transfer_rule]: "((=) ===> pcr_word) int of_nat" by transfer_prover lemma [transfer_rule]: "((=) ===> pcr_word) (\k. k) of_int" proof - have "((=) ===> pcr_word) of_int of_int" by transfer_prover then show ?thesis by (simp add: id_def) qed end +lemma abs_word_eq: + "abs_word = of_int" + by (rule ext) (transfer, rule) + context semiring_1 begin lift_definition unsigned :: "'b::len0 word \ 'a" is "of_nat \ nat \ take_bit LENGTH('b)" by simp lemma unsigned_0 [simp]: "unsigned 0 = 0" by transfer simp end context semiring_char_0 begin lemma word_eq_iff_unsigned: "a = b \ unsigned a = unsigned b" by safe (transfer; simp add: eq_nat_nat_iff) end instantiation word :: (len0) equal begin definition equal_word :: "'a word \ 'a word \ bool" where "equal_word a b \ (unsigned a :: int) = unsigned b" instance proof fix a b :: "'a word" show "HOL.equal a b \ a = b" using word_eq_iff_unsigned [of a b] by (auto simp add: equal_word_def) qed end context ring_1 begin lift_definition signed :: "'b::len word \ 'a" is "of_int \ signed_take_bit (LENGTH('b) - 1)" by (simp add: signed_take_bit_eq_iff_take_bit_eq [symmetric]) lemma signed_0 [simp]: "signed 0 = 0" by transfer simp end lemma unsigned_of_nat [simp]: "unsigned (of_nat n :: 'a word) = take_bit LENGTH('a::len) n" by transfer (simp add: nat_eq_iff take_bit_eq_mod zmod_int) lemma of_nat_unsigned [simp]: "of_nat (unsigned a) = a" by transfer simp lemma of_int_unsigned [simp]: "of_int (unsigned a) = a" by transfer simp +lemma unsigned_nat_less: + \unsigned a < (2 ^ LENGTH('a) :: nat)\ for a :: \'a::len0 word\ + by transfer (simp add: take_bit_eq_mod) + +lemma unsigned_int_less: + \unsigned a < (2 ^ LENGTH('a) :: int)\ for a :: \'a::len0 word\ + by transfer (simp add: take_bit_eq_mod) + context ring_char_0 begin lemma word_eq_iff_signed: "a = b \ signed a = signed b" by safe (transfer; auto simp add: signed_take_bit_eq_iff_take_bit_eq) end lemma signed_of_int [simp]: "signed (of_int k :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) k" by transfer simp lemma of_int_signed [simp]: "of_int (signed a) = a" by transfer (simp add: signed_take_bit_eq_take_bit take_bit_eq_mod mod_simps) subsubsection \Properties\ +lemma length_cases: + obtains (triv) "LENGTH('a::len) = 1" "take_bit LENGTH('a) 2 = (0 :: int)" + | (take_bit_2) "take_bit LENGTH('a) 2 = (2 :: int)" +proof (cases "LENGTH('a) \ 2") + case False + then have "LENGTH('a) = 1" + by (auto simp add: not_le dest: less_2_cases) + then have "take_bit LENGTH('a) 2 = (0 :: int)" + by simp + with \LENGTH('a) = 1\ triv show ?thesis + by simp +next + case True + then obtain n where "LENGTH('a) = Suc (Suc n)" + by (auto dest: le_Suc_ex) + then have "take_bit LENGTH('a) 2 = (2 :: int)" + by simp + with take_bit_2 show ?thesis + by simp +qed + subsubsection \Division\ instantiation word :: (len0) modulo begin lift_definition divide_word :: "'a word \ 'a word \ 'a word" is "\a b. take_bit LENGTH('a) a div take_bit LENGTH('a) b" by simp lift_definition modulo_word :: "'a word \ 'a word \ 'a word" is "\a b. take_bit LENGTH('a) a mod take_bit LENGTH('a) b" by simp instance .. end +lemma zero_word_div_eq [simp]: + \0 div a = 0\ for a :: \'a::len0 word\ + by transfer simp + +lemma div_zero_word_eq [simp]: + \a div 0 = 0\ for a :: \'a::len0 word\ + by transfer simp + context includes lifting_syntax begin lemma [transfer_rule]: "(pcr_word ===> (\)) even ((dvd) 2 :: 'a::len word \ bool)" proof - have even_word_unfold: "even k \ (\l. take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l))" (is "?P \ ?Q") for k :: int proof assume ?P then show ?Q by auto next assume ?Q then obtain l where "take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l)" .. then have "even (take_bit LENGTH('a) k)" by simp then show ?P by simp qed show ?thesis by (simp only: even_word_unfold [abs_def] dvd_def [where ?'a = "'a word", abs_def]) transfer_prover qed end instance word :: (len) semiring_modulo proof show "a div b * b + a mod b = a" for a b :: "'a word" proof transfer fix k l :: int define r :: int where "r = 2 ^ LENGTH('a)" then have r: "take_bit LENGTH('a) k = k mod r" for k by (simp add: take_bit_eq_mod) have "k mod r = ((k mod r) div (l mod r) * (l mod r) + (k mod r) mod (l mod r)) mod r" by (simp add: div_mult_mod_eq) also have "... = (((k mod r) div (l mod r) * (l mod r)) mod r + (k mod r) mod (l mod r)) mod r" by (simp add: mod_add_left_eq) also have "... = (((k mod r) div (l mod r) * l) mod r + (k mod r) mod (l mod r)) mod r" by (simp add: mod_mult_right_eq) finally have "k mod r = ((k mod r) div (l mod r) * l + (k mod r) mod (l mod r)) mod r" by (simp add: mod_simps) with r show "take_bit LENGTH('a) (take_bit LENGTH('a) k div take_bit LENGTH('a) l * l + take_bit LENGTH('a) k mod take_bit LENGTH('a) l) = take_bit LENGTH('a) k" by simp qed qed instance word :: (len) semiring_parity proof show "\ 2 dvd (1::'a word)" by transfer simp - consider (triv) "LENGTH('a) = 1" "take_bit LENGTH('a) 2 = (0 :: int)" - | (take_bit_2) "take_bit LENGTH('a) 2 = (2 :: int)" - proof (cases "LENGTH('a) \ 2") - case False - then have "LENGTH('a) = 1" - by (auto simp add: not_le dest: less_2_cases) - then have "take_bit LENGTH('a) 2 = (0 :: int)" - by simp - with \LENGTH('a) = 1\ triv show ?thesis - by simp - next - case True - then obtain n where "LENGTH('a) = Suc (Suc n)" - by (auto dest: le_Suc_ex) - then have "take_bit LENGTH('a) 2 = (2 :: int)" - by simp - with take_bit_2 show ?thesis - by simp - qed - note * = this show even_iff_mod_2_eq_0: "2 dvd a \ a mod 2 = 0" for a :: "'a word" - by (transfer; cases rule: *) (simp_all add: mod_2_eq_odd) + by (transfer; cases rule: length_cases [where ?'a = 'a]) (simp_all add: mod_2_eq_odd) show "\ 2 dvd a \ a mod 2 = 1" for a :: "'a word" - by (transfer; cases rule: *) (simp_all add: mod_2_eq_odd) + by (transfer; cases rule: length_cases [where ?'a = 'a]) (simp_all add: mod_2_eq_odd) qed subsubsection \Orderings\ instantiation word :: (len0) linorder begin lift_definition less_eq_word :: "'a word \ 'a word \ bool" is "\a b. take_bit LENGTH('a) a \ take_bit LENGTH('a) b" by simp lift_definition less_word :: "'a word \ 'a word \ bool" is "\a b. take_bit LENGTH('a) a < take_bit LENGTH('a) b" by simp instance by standard (transfer; auto)+ end context linordered_semidom begin lemma word_less_eq_iff_unsigned: "a \ b \ unsigned a \ unsigned b" by (transfer fixing: less_eq) (simp add: nat_le_eq_zle) lemma word_less_iff_unsigned: "a < b \ unsigned a < unsigned b" by (transfer fixing: less) (auto dest: preorder_class.le_less_trans [OF take_bit_nonnegative]) end +lemma word_greater_zero_iff: + \a > 0 \ a \ 0\ for a :: \'a::len0 word\ + by transfer (simp add: less_le) + +lemma of_nat_word_eq_iff: + \of_nat m = (of_nat n :: 'a::len word) \ take_bit LENGTH('a) m = take_bit LENGTH('a) n\ + by transfer (simp add: take_bit_of_nat) + +lemma of_nat_word_less_eq_iff: + \of_nat m \ (of_nat n :: 'a::len word) \ take_bit LENGTH('a) m \ take_bit LENGTH('a) n\ + by transfer (simp add: take_bit_of_nat) + +lemma of_nat_word_less_iff: + \of_nat m < (of_nat n :: 'a::len word) \ take_bit LENGTH('a) m < take_bit LENGTH('a) n\ + by transfer (simp add: take_bit_of_nat) + +lemma of_nat_word_eq_0_iff: + \of_nat n = (0 :: 'a::len word) \ 2 ^ LENGTH('a) dvd n\ + using of_nat_word_eq_iff [where ?'a = 'a, of n 0] by (simp add: take_bit_eq_0_iff) + +lemma of_int_word_eq_iff: + \of_int k = (of_int l :: 'a::len word) \ take_bit LENGTH('a) k = take_bit LENGTH('a) l\ + by transfer rule + +lemma of_int_word_less_eq_iff: + \of_int k \ (of_int l :: 'a::len word) \ take_bit LENGTH('a) k \ take_bit LENGTH('a) l\ + by transfer rule + +lemma of_int_word_less_iff: + \of_int k < (of_int l :: 'a::len word) \ take_bit LENGTH('a) k < take_bit LENGTH('a) l\ + by transfer rule + +lemma of_int_word_eq_0_iff: + \of_int k = (0 :: 'a::len word) \ 2 ^ LENGTH('a) dvd k\ + using of_int_word_eq_iff [where ?'a = 'a, of k 0] by (simp add: take_bit_eq_0_iff) + + subsection \Bit operation on \<^typ>\'a word\\ context unique_euclidean_semiring_with_nat begin primrec n_bits_of :: "nat \ 'a \ bool list" where "n_bits_of 0 a = []" | "n_bits_of (Suc n) a = odd a # n_bits_of n (a div 2)" lemma n_bits_of_eq_iff: "n_bits_of n a = n_bits_of n b \ take_bit n a = take_bit n b" apply (induction n arbitrary: a b) - apply auto - apply (metis local.dvd_add_times_triv_left_iff local.dvd_triv_right local.odd_one) - apply (metis local.dvd_add_times_triv_left_iff local.dvd_triv_right local.odd_one) + apply (auto elim!: evenE oddE) + apply (metis dvd_triv_right even_plus_one_iff) + apply (metis dvd_triv_right even_plus_one_iff) done lemma take_n_bits_of [simp]: "take m (n_bits_of n a) = n_bits_of (min m n) a" proof - define q and v and w where "q = min m n" and "v = m - q" and "w = n - q" then have "v = 0 \ w = 0" by auto then have "take (q + v) (n_bits_of (q + w) a) = n_bits_of q a" by (induction q arbitrary: a) auto with q_def v_def w_def show ?thesis by simp qed lemma unsigned_of_bits_n_bits_of [simp]: "unsigned_of_bits (n_bits_of n a) = take_bit n a" by (induction n arbitrary: a) (simp_all add: ac_simps) end lemma unsigned_of_bits_eq_of_bits: "unsigned_of_bits bs = (of_bits (bs @ [False]) :: int)" by (simp add: of_bits_int_def) instantiation word :: (len) bit_representation begin lift_definition bits_of_word :: "'a word \ bool list" is "n_bits_of LENGTH('a)" by (simp add: n_bits_of_eq_iff) lift_definition of_bits_word :: "bool list \ 'a word" is unsigned_of_bits . instance proof fix a :: "'a word" show "of_bits (bits_of a) = a" by transfer simp qed end lemma take_bit_complement_iff: "take_bit n (complement k) = take_bit n (complement l) \ take_bit n k = take_bit n l" for k l :: int by (simp add: take_bit_eq_mod mod_eq_dvd_iff dvd_diff_commute) lemma take_bit_not_iff: "take_bit n (NOT k) = take_bit n (NOT l) \ take_bit n k = take_bit n l" for k l :: int by (simp add: not_int_def take_bit_complement_iff) lemma n_bits_of_not: "n_bits_of n (NOT k) = map Not (n_bits_of n k)" for k :: int by (induction n arbitrary: k) (simp_all add: not_div_2) lemma take_bit_and [simp]: "take_bit n (k AND l) = take_bit n k AND take_bit n l" for k l :: int apply (induction n arbitrary: k l) apply simp apply (subst and_int.rec) apply (subst (2) and_int.rec) apply simp done lemma take_bit_or [simp]: "take_bit n (k OR l) = take_bit n k OR take_bit n l" for k l :: int apply (induction n arbitrary: k l) apply simp apply (subst or_int.rec) apply (subst (2) or_int.rec) apply simp done lemma take_bit_xor [simp]: "take_bit n (k XOR l) = take_bit n k XOR take_bit n l" for k l :: int apply (induction n arbitrary: k l) apply simp apply (subst xor_int.rec) apply (subst (2) xor_int.rec) apply simp done instantiation word :: (len) bit_operations begin lift_definition not_word :: "'a word \ 'a word" is not by (simp add: take_bit_not_iff) lift_definition and_word :: "'a word \ 'a word \ 'a word" is "and" by simp lift_definition or_word :: "'a word \ 'a word \ 'a word" is or by simp lift_definition xor_word :: "'a word \ 'a word \ 'a word" is xor by simp lift_definition shift_left_word :: "'a word \ nat \ 'a word" is shift_left proof - show "take_bit LENGTH('a) (k << n) = take_bit LENGTH('a) (l << n)" if "take_bit LENGTH('a) k = take_bit LENGTH('a) l" for k l :: int and n :: nat proof - from that have "take_bit (LENGTH('a) - n) (take_bit LENGTH('a) k) = take_bit (LENGTH('a) - n) (take_bit LENGTH('a) l)" by simp moreover have "min (LENGTH('a) - n) LENGTH('a) = LENGTH('a) - n" by simp ultimately show ?thesis by (simp add: take_bit_push_bit) qed qed lift_definition shift_right_word :: "'a word \ nat \ 'a word" is "\k n. drop_bit n (take_bit LENGTH('a) k)" by simp instance proof show "semilattice ((AND) :: 'a word \ _)" by standard (transfer; simp add: ac_simps)+ show "semilattice ((OR) :: 'a word \ _)" by standard (transfer; simp add: ac_simps)+ show "abel_semigroup ((XOR) :: 'a word \ _)" by standard (transfer; simp add: ac_simps)+ show "not = (of_bits \ map Not \ bits_of :: 'a word \ 'a word)" proof fix a :: "'a word" have "NOT a = of_bits (map Not (bits_of a))" by transfer (simp flip: unsigned_of_bits_take n_bits_of_not add: take_map) then show "NOT a = (of_bits \ map Not \ bits_of) a" by simp qed show "of_bits bs AND of_bits cs = (of_bits (map2 (\) bs cs) :: 'a word)" if "length bs = length cs" for bs cs using that apply transfer apply (simp only: unsigned_of_bits_eq_of_bits) apply (subst and_eq) apply simp_all done show "of_bits bs OR of_bits cs = (of_bits (map2 (\) bs cs) :: 'a word)" if "length bs = length cs" for bs cs using that apply transfer apply (simp only: unsigned_of_bits_eq_of_bits) apply (subst or_eq) apply simp_all done show "of_bits bs XOR of_bits cs = (of_bits (map2 (\) bs cs) :: 'a word)" if "length bs = length cs" for bs cs using that apply transfer apply (simp only: unsigned_of_bits_eq_of_bits) apply (subst xor_eq) apply simp_all done show "a << n = of_bits (replicate n False @ bits_of a)" for a :: "'a word" and n :: nat by transfer (simp add: push_bit_take_bit) show "a >> n = of_bits (drop n (bits_of a))" if "n < length (bits_of a)" for a :: "'a word" and n :: nat using that by transfer simp qed + +subsection \Bit structure on \<^typ>\'a word\\ + +lemma word_bit_induct [case_names zero even odd]: + \P a\ if word_zero: \P 0\ + and word_even: \\a. P a \ 0 < a \ a < 2 ^ (LENGTH('a) - 1) \ P (2 * a)\ + and word_odd: \\a. P a \ a < 2 ^ (LENGTH('a) - 1) \ P (1 + 2 * a)\ + for P and a :: \'a::len word\ +proof - + define m :: nat where \m = LENGTH('a) - 1\ + then have l: \LENGTH('a) = Suc m\ + by simp + define n :: nat where \n = unsigned a\ + then have \n < 2 ^ LENGTH('a)\ + by (simp add: unsigned_nat_less) + then have \n < 2 * 2 ^ m\ + by (simp add: l) + then have \P (of_nat n)\ + proof (induction n rule: nat_bit_induct) + case zero + show ?case + by simp (rule word_zero) + next + case (even n) + then have \n < 2 ^ m\ + by simp + with even.IH have \P (of_nat n)\ + by simp + moreover from \n < 2 ^ m\ even.hyps have \0 < (of_nat n :: 'a word)\ + by (auto simp add: word_greater_zero_iff of_nat_word_eq_0_iff l) + moreover from \n < 2 ^ m\ have \(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - 1)\ + using of_nat_word_less_iff [where ?'a = 'a, of n \2 ^ m\] + by (cases \m = 0\) (simp_all add: not_less take_bit_eq_self ac_simps l) + ultimately have \P (2 * of_nat n)\ + by (rule word_even) + then show ?case + by simp + next + case (odd n) + then have \Suc n \ 2 ^ m\ + by simp + with odd.IH have \P (of_nat n)\ + by simp + moreover from \Suc n \ 2 ^ m\ have \(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - 1)\ + using of_nat_word_less_iff [where ?'a = 'a, of n \2 ^ m\] + by (cases \m = 0\) (simp_all add: not_less take_bit_eq_self ac_simps l) + ultimately have \P (1 + 2 * of_nat n)\ + by (rule word_odd) + then show ?case + by simp + qed + then show ?thesis + by (simp add: n_def) +qed + end global_interpretation bit_word: boolean_algebra "(AND)" "(OR)" NOT 0 "- 1 :: 'a::len word" rewrites "bit_word.xor = ((XOR) :: 'a word \ _)" proof - interpret bit_word: boolean_algebra "(AND)" "(OR)" NOT 0 "- 1 :: 'a word" proof show "a AND (b OR c) = a AND b OR a AND c" for a b c :: "'a word" by transfer (simp add: bit_int.conj_disj_distrib) show "a OR b AND c = (a OR b) AND (a OR c)" for a b c :: "'a word" by transfer (simp add: bit_int.disj_conj_distrib) show "a AND NOT a = 0" for a :: "'a word" by transfer simp show "a OR NOT a = - 1" for a :: "'a word" by transfer simp qed (transfer; simp)+ show "boolean_algebra (AND) (OR) NOT 0 (- 1 :: 'a word)" by (fact bit_word.boolean_algebra_axioms) show "bit_word.xor = ((XOR) :: 'a word \ _)" proof (rule ext)+ fix a b :: "'a word" have "a XOR b = a AND NOT b OR NOT a AND b" by transfer (simp add: bit_int.xor_def) then show "bit_word.xor a b = a XOR b" by (simp add: bit_word.xor_def) qed qed end