# HG changeset patch # User haftmann # Date 1555567614 0 # Node ID 3ea80c9500234022cede5272798ad1cff5ea6c88 # Parent ca9dfa7ee3bd1a60eb11cf1cf0449311bf3f9915 incorporated various material from the AFP into the distribution diff -r ca9dfa7ee3bd -r 3ea80c950023 src/HOL/List.thy --- a/src/HOL/List.thy Wed Apr 17 16:57:06 2019 +0000 +++ b/src/HOL/List.thy Thu Apr 18 06:06:54 2019 +0000 @@ -1995,6 +1995,9 @@ using longest_common_prefix[of "rev xs" "rev ys"] unfolding rev_swap rev_append by (metis last_rev rev_is_Nil_conv) +lemma butlast_rev [simp]: "butlast (rev xs) = rev (tl xs)" + by (cases xs) simp_all + subsubsection \\<^const>\take\ and \<^const>\drop\\ @@ -3150,7 +3153,7 @@ lemma hd_upt[simp]: "i < j \ hd[i.. last[i..More lemmas\ + +lemma twice_conv_BIT: "2 * x = x BIT False" + by (rule bin_rl_eqI) (simp_all, simp_all add: bin_rest_def bin_last_def) + +lemma not_int_cmp_0 [simp]: + fixes i :: int shows + "0 < NOT i \ i < -1" + "0 \ NOT i \ i < 0" + "NOT i < 0 \ i \ 0" + "NOT i \ 0 \ i \ -1" +by(simp_all add: int_not_def) arith+ + +lemma bbw_ao_dist2: "(x :: int) AND (y OR z) = x AND y OR x AND z" +by(metis int_and_comm bbw_ao_dist) + +lemmas int_and_ac = bbw_lcs(1) int_and_comm int_and_assoc + +lemma int_nand_same [simp]: fixes x :: int shows "x AND NOT x = 0" +by(induct x y\"NOT x" rule: bitAND_int.induct)(subst bitAND_int.simps, clarsimp) + +lemma int_nand_same_middle: fixes x :: int shows "x AND y AND NOT x = 0" +by (metis bbw_lcs(1) int_and_0 int_nand_same) + +lemma and_xor_dist: fixes x :: int shows + "x AND (y XOR z) = (x AND y) XOR (x AND z)" +by(simp add: int_xor_def bbw_ao_dist2 bbw_ao_dist bbw_not_dist int_and_ac int_nand_same_middle) + +lemma BIT_lt0 [simp]: "x BIT b < 0 \ x < 0" +by(cases b)(auto simp add: Bit_def) + +lemma BIT_ge0 [simp]: "x BIT b \ 0 \ x \ 0" +by(cases b)(auto simp add: Bit_def) + +lemma [simp]: + shows bin_rest_lt0: "bin_rest i < 0 \ i < 0" + and bin_rest_ge_0: "bin_rest i \ 0 \ i \ 0" +by(auto simp add: bin_rest_def) + +lemma bin_rest_gt_0 [simp]: "bin_rest x > 0 \ x > 1" +by(simp add: bin_rest_def add1_zle_eq pos_imp_zdiv_pos_iff) (metis add1_zle_eq one_add_one) + +lemma int_and_lt0 [simp]: fixes x y :: int shows + "x AND y < 0 \ x < 0 \ y < 0" +by(induct x y rule: bitAND_int.induct)(subst bitAND_int.simps, simp) + +lemma int_and_ge0 [simp]: fixes x y :: int shows + "x AND y \ 0 \ x \ 0 \ y \ 0" +by (metis int_and_lt0 linorder_not_less) + +lemma int_and_1: fixes x :: int shows "x AND 1 = x mod 2" +by(subst bitAND_int.simps)(simp add: Bit_def bin_last_def zmod_minus1) + +lemma int_1_and: fixes x :: int shows "1 AND x = x mod 2" +by(subst int_and_comm)(simp add: int_and_1) + +lemma int_or_lt0 [simp]: fixes x y :: int shows + "x OR y < 0 \ x < 0 \ y < 0" +by(simp add: int_or_def) + +lemma int_xor_lt0 [simp]: fixes x y :: int shows + "x XOR y < 0 \ ((x < 0) \ (y < 0))" +by(auto simp add: int_xor_def) + +lemma int_xor_ge0 [simp]: fixes x y :: int shows + "x XOR y \ 0 \ ((x \ 0) \ (y \ 0))" +by (metis int_xor_lt0 linorder_not_le) + +lemma bin_last_conv_AND: + "bin_last i \ i AND 1 \ 0" +proof - + obtain x b where "i = x BIT b" by(cases i rule: bin_exhaust) + hence "i AND 1 = 0 BIT b" + by(simp add: BIT_special_simps(2)[symmetric] del: BIT_special_simps(2)) + thus ?thesis using \i = x BIT b\ by(cases b) simp_all +qed + +lemma bitval_bin_last: + "of_bool (bin_last i) = i AND 1" +proof - + obtain x b where "i = x BIT b" by(cases i rule: bin_exhaust) + hence "i AND 1 = 0 BIT b" + by(simp add: BIT_special_simps(2)[symmetric] del: BIT_special_simps(2)) + thus ?thesis by(cases b)(simp_all add: bin_last_conv_AND) +qed + +lemma bl_to_bin_BIT: + "bl_to_bin bs BIT b = bl_to_bin (bs @ [b])" +by(simp add: bl_to_bin_append) + +lemma bin_last_bl_to_bin: "bin_last (bl_to_bin bs) \ bs \ [] \ last bs" +by(cases "bs = []")(auto simp add: bl_to_bin_def last_bin_last'[where w=0]) + +lemma bin_rest_bl_to_bin: "bin_rest (bl_to_bin bs) = bl_to_bin (butlast bs)" +by(cases "bs = []")(simp_all add: bl_to_bin_def butlast_rest_bl2bin_aux) + +lemma bin_nth_numeral_unfold: + "bin_nth (numeral (num.Bit0 x)) n \ n > 0 \ bin_nth (numeral x) (n - 1)" + "bin_nth (numeral (num.Bit1 x)) n \ (n > 0 \ bin_nth (numeral x) (n - 1))" +by(case_tac [!] n) simp_all + +lemma bin_sign_and: + "bin_sign (i AND j) = - (bin_sign i * bin_sign j)" +by(simp add: bin_sign_def) + +lemma minus_BIT_0: fixes x y :: int shows "x BIT b - y BIT False = (x - y) BIT b" +by(simp add: Bit_def) + +lemma int_not_neg_numeral: "NOT (- numeral n) = (Num.sub n num.One :: int)" +by(simp add: int_not_def) + +lemma sub_inc_One: "Num.sub (Num.inc n) num.One = numeral n" +by (metis add_diff_cancel diff_minus_eq_add diff_numeral_special(2) diff_numeral_special(6)) + +lemma inc_BitM: "Num.inc (Num.BitM n) = num.Bit0 n" +by(simp add: BitM_plus_one[symmetric] add_One) + +lemma int_neg_numeral_pOne_conv_not: "- numeral (n + num.One) = (NOT (numeral n) :: int)" +by(simp add: int_not_def) + +lemma int_lsb_BIT [simp]: fixes x :: int shows + "lsb (x BIT b) \ b" +by(simp add: lsb_int_def) + +lemma bin_last_conv_lsb: "bin_last = lsb" +by(clarsimp simp add: lsb_int_def fun_eq_iff) + +lemma int_lsb_numeral [simp]: + "lsb (0 :: int) = False" + "lsb (1 :: int) = True" + "lsb (Numeral1 :: int) = True" + "lsb (- 1 :: int) = True" + "lsb (- Numeral1 :: int) = True" + "lsb (numeral (num.Bit0 w) :: int) = False" + "lsb (numeral (num.Bit1 w) :: int) = True" + "lsb (- numeral (num.Bit0 w) :: int) = False" + "lsb (- numeral (num.Bit1 w) :: int) = True" +by(simp_all add: lsb_int_def) + +lemma int_set_bit_0 [simp]: fixes x :: int shows + "set_bit x 0 b = bin_rest x BIT b" +by(auto simp add: set_bit_int_def intro: bin_rl_eqI) + +lemma int_set_bit_Suc: fixes x :: int shows + "set_bit x (Suc n) b = set_bit (bin_rest x) n b BIT bin_last x" +by(auto simp add: set_bit_int_def twice_conv_BIT intro: bin_rl_eqI) + +lemma bin_last_set_bit: + "bin_last (set_bit x n b) = (if n > 0 then bin_last x else b)" +by(cases n)(simp_all add: int_set_bit_Suc) + +lemma bin_rest_set_bit: + "bin_rest (set_bit x n b) = (if n > 0 then set_bit (bin_rest x) (n - 1) b else bin_rest x)" +by(cases n)(simp_all add: int_set_bit_Suc) + +lemma int_set_bit_numeral: fixes x :: int shows + "set_bit x (numeral w) b = set_bit (bin_rest x) (pred_numeral w) b BIT bin_last x" +by(simp add: set_bit_int_def) + +lemmas int_set_bit_numerals [simp] = + int_set_bit_numeral[where x="numeral w'"] + int_set_bit_numeral[where x="- numeral w'"] + int_set_bit_numeral[where x="Numeral1"] + int_set_bit_numeral[where x="1"] + int_set_bit_numeral[where x="0"] + int_set_bit_Suc[where x="numeral w'"] + int_set_bit_Suc[where x="- numeral w'"] + int_set_bit_Suc[where x="Numeral1"] + int_set_bit_Suc[where x="1"] + int_set_bit_Suc[where x="0"] + for w' + +lemma int_shiftl_BIT: fixes x :: int + shows int_shiftl0 [simp]: "x << 0 = x" + and int_shiftl_Suc [simp]: "x << Suc n = (x << n) BIT False" +by(auto simp add: shiftl_int_def Bit_def) + +lemma int_0_shiftl [simp]: "0 << n = (0 :: int)" +by(induct n) simp_all + +lemma bin_last_shiftl: "bin_last (x << n) \ n = 0 \ bin_last x" +by(cases n)(simp_all) + +lemma bin_rest_shiftl: "bin_rest (x << n) = (if n > 0 then x << (n - 1) else bin_rest x)" +by(cases n)(simp_all) + +lemma bin_nth_shiftl [simp]: "bin_nth (x << n) m \ n \ m \ bin_nth x (m - n)" +proof(induct n arbitrary: x m) + case (Suc n) + thus ?case by(cases m) simp_all +qed simp + +lemma int_shiftr_BIT [simp]: fixes x :: int + shows int_shiftr0: "x >> 0 = x" + and int_shiftr_Suc: "x BIT b >> Suc n = x >> n" +proof - + show "x >> 0 = x" by (simp add: shiftr_int_def) + show "x BIT b >> Suc n = x >> n" by (cases b) + (simp_all add: shiftr_int_def Bit_def add.commute pos_zdiv_mult_2) +qed + +lemma bin_last_shiftr: "bin_last (x >> n) \ x !! n" +proof(induct n arbitrary: x) + case 0 thus ?case by simp +next + case (Suc n) + thus ?case by(cases x rule: bin_exhaust) simp +qed + +lemma bin_rest_shiftr [simp]: "bin_rest (x >> n) = x >> Suc n" +proof(induct n arbitrary: x) + case 0 + thus ?case by(cases x rule: bin_exhaust) auto +next + case (Suc n) + thus ?case by(cases x rule: bin_exhaust) auto +qed + +lemma bin_nth_shiftr [simp]: "bin_nth (x >> n) m = bin_nth x (n + m)" +proof(induct n arbitrary: x m) + case (Suc n) + thus ?case by(cases x rule: bin_exhaust) simp_all +qed simp + +lemma bin_nth_conv_AND: + fixes x :: int shows + "bin_nth x n \ x AND (1 << n) \ 0" +proof(induct n arbitrary: x) + case 0 + thus ?case by(simp add: int_and_1 bin_last_def) +next + case (Suc n) + thus ?case by(cases x rule: bin_exhaust)(simp_all add: bin_nth_ops Bit_eq_0_iff) +qed + +lemma int_shiftl_numeral [simp]: + "(numeral w :: int) << numeral w' = numeral (num.Bit0 w) << pred_numeral w'" + "(- numeral w :: int) << numeral w' = - numeral (num.Bit0 w) << pred_numeral w'" +by(simp_all add: numeral_eq_Suc Bit_def shiftl_int_def) + (metis add_One mult_inc semiring_norm(11) semiring_norm(13) semiring_norm(2) semiring_norm(6) semiring_norm(87))+ + +lemma int_shiftl_One_numeral [simp]: "(1 :: int) << numeral w = 2 << pred_numeral w" +by(metis int_shiftl_numeral numeral_One) + +lemma shiftl_ge_0 [simp]: fixes i :: int shows "i << n \ 0 \ i \ 0" +by(induct n) simp_all + +lemma shiftl_lt_0 [simp]: fixes i :: int shows "i << n < 0 \ i < 0" +by (metis not_le shiftl_ge_0) + +lemma int_shiftl_test_bit: "(n << i :: int) !! m \ m \ i \ n !! (m - i)" +proof(induction i) + case (Suc n) + thus ?case by(cases m) simp_all +qed simp + +lemma int_0shiftr [simp]: "(0 :: int) >> x = 0" +by(simp add: shiftr_int_def) + +lemma int_minus1_shiftr [simp]: "(-1 :: int) >> x = -1" +by(simp add: shiftr_int_def div_eq_minus1) + +lemma int_shiftr_ge_0 [simp]: fixes i :: int shows "i >> n \ 0 \ i \ 0" +proof(induct n arbitrary: i) + case (Suc n) + thus ?case by(cases i rule: bin_exhaust) simp_all +qed simp + +lemma int_shiftr_lt_0 [simp]: fixes i :: int shows "i >> n < 0 \ i < 0" +by (metis int_shiftr_ge_0 not_less) + +lemma uminus_Bit_eq: + "- k BIT b = (- k - of_bool b) BIT b" + by (cases b) (simp_all add: Bit_def) + +lemma int_shiftr_numeral [simp]: + "(1 :: int) >> numeral w' = 0" + "(numeral num.One :: int) >> numeral w' = 0" + "(numeral (num.Bit0 w) :: int) >> numeral w' = numeral w >> pred_numeral w'" + "(numeral (num.Bit1 w) :: int) >> numeral w' = numeral w >> pred_numeral w'" + "(- numeral (num.Bit0 w) :: int) >> numeral w' = - numeral w >> pred_numeral w'" + "(- numeral (num.Bit1 w) :: int) >> numeral w' = - numeral (Num.inc w) >> pred_numeral w'" + by (simp_all only: numeral_One expand_BIT numeral_eq_Suc int_shiftr_Suc BIT_special_simps(2)[symmetric] int_0shiftr add_One uminus_Bit_eq) + (simp_all add: add_One) + +lemma int_shiftr_numeral_Suc0 [simp]: + "(1 :: int) >> Suc 0 = 0" + "(numeral num.One :: int) >> Suc 0 = 0" + "(numeral (num.Bit0 w) :: int) >> Suc 0 = numeral w" + "(numeral (num.Bit1 w) :: int) >> Suc 0 = numeral w" + "(- numeral (num.Bit0 w) :: int) >> Suc 0 = - numeral w" + "(- numeral (num.Bit1 w) :: int) >> Suc 0 = - numeral (Num.inc w)" +by(simp_all only: One_nat_def[symmetric] numeral_One[symmetric] int_shiftr_numeral pred_numeral_simps int_shiftr0) + +lemma bin_nth_minus_p2: + assumes sign: "bin_sign x = 0" + and y: "y = 1 << n" + and m: "m < n" + and x: "x < y" + shows "bin_nth (x - y) m = bin_nth x m" +using sign m x unfolding y +proof(induction m arbitrary: x y n) + case 0 + thus ?case + by(simp add: bin_last_def shiftl_int_def) (metis (hide_lams, no_types) mod_diff_right_eq mod_self neq0_conv numeral_One power_eq_0_iff power_mod diff_zero zero_neq_numeral) +next + case (Suc m) + from \Suc m < n\ obtain n' where [simp]: "n = Suc n'" by(cases n) auto + obtain x' b where [simp]: "x = x' BIT b" by(cases x rule: bin_exhaust) + from \bin_sign x = 0\ have "bin_sign x' = 0" by simp + moreover from \x < 1 << n\ have "x' < 1 << n'" + by(cases b)(simp_all add: Bit_def shiftl_int_def) + moreover have "(2 * x' + of_bool b - 2 * 2 ^ n') div 2 = x' + (- (2 ^ n') + of_bool b div 2)" + by(simp only: add_diff_eq[symmetric] add.commute div_mult_self2[OF zero_neq_numeral[symmetric]]) + ultimately show ?case using Suc.IH[of x' n'] Suc.prems + by(cases b)(simp_all add: Bit_def bin_rest_def shiftl_int_def) +qed + +lemma bin_clr_conv_NAND: + "bin_sc n False i = i AND NOT (1 << n)" +by(induct n arbitrary: i)(auto intro: bin_rl_eqI) + +lemma bin_set_conv_OR: + "bin_sc n True i = i OR (1 << n)" +by(induct n arbitrary: i)(auto intro: bin_rl_eqI) + +lemma int_set_bits_K_True [simp]: "(BITS _. True) = (-1 :: int)" +by(auto simp add: set_bits_int_def bin_last_bl_to_bin) + +lemma int_set_bits_K_False [simp]: "(BITS _. False) = (0 :: int)" +by(auto simp add: set_bits_int_def) + +lemma msb_conv_bin_sign: "msb x \ bin_sign x = -1" +by(simp add: bin_sign_def not_le msb_int_def) + +lemma msb_BIT [simp]: "msb (x BIT b) = msb x" +by(simp add: msb_int_def) + +lemma msb_bin_rest [simp]: "msb (bin_rest x) = msb x" +by(simp add: msb_int_def) + +lemma int_msb_and [simp]: "msb ((x :: int) AND y) \ msb x \ msb y" +by(simp add: msb_int_def) + +lemma int_msb_or [simp]: "msb ((x :: int) OR y) \ msb x \ msb y" +by(simp add: msb_int_def) + +lemma int_msb_xor [simp]: "msb ((x :: int) XOR y) \ msb x \ msb y" +by(simp add: msb_int_def) + +lemma int_msb_not [simp]: "msb (NOT (x :: int)) \ \ msb x" +by(simp add: msb_int_def not_less) + +lemma msb_shiftl [simp]: "msb ((x :: int) << n) \ msb x" +by(simp add: msb_int_def) + +lemma msb_shiftr [simp]: "msb ((x :: int) >> r) \ msb x" +by(simp add: msb_int_def) + +lemma msb_bin_sc [simp]: "msb (bin_sc n b x) \ msb x" +by(simp add: msb_conv_bin_sign) + +lemma msb_set_bit [simp]: "msb (set_bit (x :: int) n b) \ msb x" +by(simp add: msb_conv_bin_sign set_bit_int_def) + +lemma msb_0 [simp]: "msb (0 :: int) = False" +by(simp add: msb_int_def) + +lemma msb_1 [simp]: "msb (1 :: int) = False" +by(simp add: msb_int_def) + +lemma msb_numeral [simp]: + "msb (numeral n :: int) = False" + "msb (- numeral n :: int) = True" +by(simp_all add: msb_int_def) + end diff -r ca9dfa7ee3bd -r 3ea80c950023 src/HOL/Word/Word.thy --- a/src/HOL/Word/Word.thy Wed Apr 17 16:57:06 2019 +0000 +++ b/src/HOL/Word/Word.thy Thu Apr 18 06:06:54 2019 +0000 @@ -96,7 +96,7 @@ instantiation word :: (len0) size begin -definition word_size: "size (w :: 'a word) = len_of TYPE('a)" +definition word_size: "size (w :: 'a word) = LENGTH('a)" instance .. @@ -1883,8 +1883,7 @@ for x y :: "'a::len word" by (simp add: uint_nat unat_mod zmod_int) - -subsection \Definition of \unat_arith\ tactic\ +text \Definition of \unat_arith\ tactic\ lemma unat_split: "P (unat x) \ (\n. of_nat n = x \ n < 2^len_of TYPE('a) \ P n)" for x :: "'a::len word" @@ -2065,16 +2064,20 @@ apply (rule bl_to_bin_lt2p) done +lemma unatSuc: "1 + n \ 0 \ unat (1 + n) = Suc (unat n)" + for n :: "'a::len word" + by unat_arith + subsection \Cardinality, finiteness of set of words\ instance word :: (len0) finite by standard (simp add: type_definition.univ [OF type_definition_word]) -lemma card_word: "CARD('a::len0 word) = 2 ^ len_of TYPE('a)" +lemma card_word: "CARD('a word) = 2 ^ LENGTH('a::len0)" by (simp add: type_definition.card [OF type_definition_word] nat_power_eq) -lemma card_word_size: "card (UNIV :: 'a word set) = (2 ^ size x)" +lemma card_word_size: "CARD('a word) = 2 ^ size x" for x :: "'a::len0 word" unfolding word_size by (rule card_word) @@ -2583,6 +2586,10 @@ apply simp done +lemma set_bit_beyond: + "size x \ n \ set_bit x n b = x" for x :: "'a :: len0 word" + by (auto intro: word_eqI simp add: test_bit_set_gen word_size) + subsection \Shifting, Rotating, and Splitting Words\ @@ -2794,8 +2801,6 @@ apply (unfold word_reverse_def) apply (rule word_bl.Rep_inverse' [symmetric]) apply (simp add: bl_shiftl1' bl_shiftr1' word_bl.Abs_inverse) - apply (cases "to_bl w") - apply auto done lemma shiftl_rev: "shiftl w n = word_reverse (shiftr (word_reverse w) n)" @@ -3162,6 +3167,9 @@ using word_of_int_Ex [where x=x] by (auto simp: and_mask_wi' word_of_int_power_hom word.abs_eq_iff bintrunc_mod2p mod_simps) +lemma mask_full [simp]: "mask LENGTH('a) = (- 1 :: 'a::len word)" + by (simp add: mask_def word_size shiftl_zero_size) + subsubsection \Revcast\ @@ -3411,6 +3419,10 @@ simp)+ done +lemma shiftr_zero_size: "size x \ n \ x >> n = 0" + for x :: "'a :: len0 word" + by (rule word_eqI) (auto simp add: nth_shiftr dest: test_bit_size) + subsection \Split and cat\ @@ -4493,9 +4505,63 @@ apply simp done -lemma unatSuc: "1 + n \ 0 \ unat (1 + n) = Suc (unat n)" - for n :: "'a::len word" - by unat_arith + +subsection \More\ + +lemma set_bits_K_False [simp]: + "set_bits (\_. False) = (0 :: 'a :: len0 word)" + by (rule word_eqI) (simp add: test_bit.eq_norm) + +lemma test_bit_1' [simp]: + "(1 :: 'a :: len0 word) !! n \ 0 < len_of TYPE('a) \ n = 0" + by (cases n) (simp_all only: one_word_def test_bit_wi bin_nth.simps, simp_all) + +lemma mask_0 [simp]: + "mask 0 = 0" + by (simp add: Word.mask_def) + +lemma shiftl0 [simp]: + "x << 0 = (x :: 'a :: len0 word)" + by (metis shiftl_rev shiftr_x_0 word_rev_gal) + +lemma mask_1: "mask 1 = 1" + by (simp add: mask_def) + +lemma mask_Suc_0: "mask (Suc 0) = 1" + by (simp add: mask_def) + +lemma mask_numeral: "mask (numeral n) = 2 * mask (pred_numeral n) + 1" + by (simp add: mask_def neg_numeral_class.sub_def numeral_eq_Suc numeral_pow) + +lemma bin_last_bintrunc: "bin_last (bintrunc l n) = (l > 0 \ bin_last n)" + by (cases l) simp_all + +lemma word_and_1: + "n AND 1 = (if n !! 0 then 1 else 0)" for n :: "_ word" + by transfer (rule bin_rl_eqI, simp_all add: bin_rest_trunc bin_last_bintrunc) + +lemma bintrunc_shiftl: + "bintrunc n (m << i) = bintrunc (n - i) m << i" +proof (induction i arbitrary: n) + case 0 + show ?case + by simp +next + case (Suc i) + then show ?case by (cases n) simp_all +qed + +lemma shiftl_transfer [transfer_rule]: + includes lifting_syntax + shows "(pcr_word ===> (=) ===> pcr_word) (<<) (<<)" + by (auto intro!: rel_funI word_eqI simp add: word.pcr_cr_eq cr_word_def word_size nth_shiftl) + +lemma uint_shiftl: + "uint (n << i) = bintrunc (size n) (uint n << i)" + unfolding word_size by transfer (simp add: bintrunc_shiftl) + + +subsection \Misc\ declare bin_to_bl_def [simp]