# HG changeset patch # User haftmann # Date 1572512522 0 # Node ID a7a52ba0717d60e514a56ccec43b24a7a52207eb # Parent 196b41b9b9c88ef8f08d305c42887429e360891b more lemmas diff -r 196b41b9b9c8 -r a7a52ba0717d src/HOL/Parity.thy --- a/src/HOL/Parity.thy Wed Oct 30 18:30:28 2019 -0400 +++ b/src/HOL/Parity.thy Thu Oct 31 09:02:02 2019 +0000 @@ -1093,6 +1093,10 @@ "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) diff -r 196b41b9b9c8 -r a7a52ba0717d src/HOL/ex/Word_Type.thy --- a/src/HOL/ex/Word_Type.thy Wed Oct 30 18:30:28 2019 -0400 +++ b/src/HOL/ex/Word_Type.thy Thu Oct 31 09:02:02 2019 +0000 @@ -141,6 +141,17 @@ "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\ @@ -169,6 +180,10 @@ end +lemma abs_word_eq: + "abs_word = of_int" + by (rule ext) (transfer, rule) + context semiring_1 begin @@ -230,6 +245,14 @@ "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 @@ -250,6 +273,27 @@ 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\ @@ -268,6 +312,14 @@ 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 @@ -325,32 +377,12 @@ 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 @@ -385,6 +417,43 @@ 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 @@ -398,9 +467,9 @@ 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]: @@ -573,6 +642,61 @@ 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"