diff --git a/src/HOL/Word/Bit_Comprehension.thy b/src/HOL/Word/Bit_Comprehension.thy --- a/src/HOL/Word/Bit_Comprehension.thy +++ b/src/HOL/Word/Bit_Comprehension.thy @@ -1,37 +1,37 @@ (* Title: HOL/Word/Bit_Comprehension.thy Author: Brian Huffman, PSU; Jeremy Dawson and Gerwin Klein, NICTA *) section \Comprehension syntax for bit expressions\ theory Bit_Comprehension imports Bits_Int begin class bit_comprehension = semiring_bits + fixes set_bits :: "(nat \ bool) \ 'a" (binder "BITS " 10) instantiation int :: bit_comprehension begin definition "set_bits f = (if \n. \n'\n. \ f n' then let n = LEAST n. \n'\n. \ f n' in horner_sum of_bool 2 (map f [0..n. \n'\n. f n' then let n = LEAST n. \n'\n. f n' in signed_take_bit n (horner_sum of_bool 2 (map f [0..Bit values as reversed lists of bools\ theory Bit_Lists imports Bits_Int begin subsection \Implicit augmentation of list prefixes\ primrec takefill :: "'a \ nat \ 'a list \ 'a list" where Z: "takefill fill 0 xs = []" | Suc: "takefill fill (Suc n) xs = (case xs of [] \ fill # takefill fill n xs | y # ys \ y # takefill fill n ys)" lemma nth_takefill: "m < n \ takefill fill n l ! m = (if m < length l then l ! m else fill)" apply (induct n arbitrary: m l) apply clarsimp apply clarsimp apply (case_tac m) apply (simp split: list.split) apply (simp split: list.split) done lemma takefill_alt: "takefill fill n l = take n l @ replicate (n - length l) fill" by (induct n arbitrary: l) (auto split: list.split) lemma takefill_replicate [simp]: "takefill fill n (replicate m fill) = replicate n fill" by (simp add: takefill_alt replicate_add [symmetric]) lemma takefill_le': "n = m + k \ takefill x m (takefill x n l) = takefill x m l" by (induct m arbitrary: l n) (auto split: list.split) lemma length_takefill [simp]: "length (takefill fill n l) = n" by (simp add: takefill_alt) lemma take_takefill': "n = k + m \ take k (takefill fill n w) = takefill fill k w" by (induct k arbitrary: w n) (auto split: list.split) lemma drop_takefill: "drop k (takefill fill (m + k) w) = takefill fill m (drop k w)" by (induct k arbitrary: w) (auto split: list.split) lemma takefill_le [simp]: "m \ n \ takefill x m (takefill x n l) = takefill x m l" by (auto simp: le_iff_add takefill_le') lemma take_takefill [simp]: "m \ n \ take m (takefill fill n w) = takefill fill m w" by (auto simp: le_iff_add take_takefill') lemma takefill_append: "takefill fill (m + length xs) (xs @ w) = xs @ (takefill fill m w)" by (induct xs) auto lemma takefill_same': "l = length xs \ takefill fill l xs = xs" by (induct xs arbitrary: l) auto lemmas takefill_same [simp] = takefill_same' [OF refl] lemma tf_rev: "n + k = m + length bl \ takefill x m (rev (takefill y n bl)) = rev (takefill y m (rev (takefill x k (rev bl))))" apply (rule nth_equalityI) apply (auto simp add: nth_takefill rev_nth) apply (rule_tac f = "\n. bl ! n" in arg_cong) apply arith done lemma takefill_minus: "0 < n \ takefill fill (Suc (n - 1)) w = takefill fill n w" by auto lemmas takefill_Suc_cases = list.cases [THEN takefill.Suc [THEN trans]] lemmas takefill_Suc_Nil = takefill_Suc_cases (1) lemmas takefill_Suc_Cons = takefill_Suc_cases (2) lemmas takefill_minus_simps = takefill_Suc_cases [THEN [2] takefill_minus [symmetric, THEN trans]] lemma takefill_numeral_Nil [simp]: "takefill fill (numeral k) [] = fill # takefill fill (pred_numeral k) []" by (simp add: numeral_eq_Suc) lemma takefill_numeral_Cons [simp]: "takefill fill (numeral k) (x # xs) = x # takefill fill (pred_numeral k) xs" by (simp add: numeral_eq_Suc) subsection \Range projection\ definition bl_of_nth :: "nat \ (nat \ 'a) \ 'a list" where "bl_of_nth n f = map f (rev [0.. rev (bl_of_nth n f) ! m = f m" by (simp add: bl_of_nth_def rev_map) lemma bl_of_nth_inj: "(\k. k < n \ f k = g k) \ bl_of_nth n f = bl_of_nth n g" by (simp add: bl_of_nth_def) lemma bl_of_nth_nth_le: "n \ length xs \ bl_of_nth n (nth (rev xs)) = drop (length xs - n) xs" apply (induct n arbitrary: xs) apply clarsimp apply clarsimp apply (rule trans [OF _ hd_Cons_tl]) apply (frule Suc_le_lessD) apply (simp add: rev_nth trans [OF drop_Suc drop_tl, symmetric]) apply (subst hd_drop_conv_nth) apply force apply simp_all apply (rule_tac f = "\n. drop n xs" in arg_cong) apply simp done lemma bl_of_nth_nth [simp]: "bl_of_nth (length xs) ((!) (rev xs)) = xs" by (simp add: bl_of_nth_nth_le) -lemma takefill_bintrunc: "takefill False n bl = rev (bin_to_bl n (bl_to_bin (rev bl)))" - apply (rule nth_equalityI) - apply simp - apply (clarsimp simp: nth_takefill rev_nth nth_bin_to_bl bin_nth_of_bl) - done - -lemma bl_bin_bl_rtf: "bin_to_bl n (bl_to_bin bl) = rev (takefill False n (rev bl))" - by (simp add: takefill_bintrunc) - subsection \More\ definition rotater1 :: "'a list \ 'a list" where "rotater1 ys = (case ys of [] \ [] | x # xs \ last ys # butlast ys)" definition rotater :: "nat \ 'a list \ 'a list" where "rotater n = rotater1 ^^ n" lemmas rotater_0' [simp] = rotater_def [where n = "0", simplified] lemma rotate1_rl': "rotater1 (l @ [a]) = a # l" by (cases l) (auto simp: rotater1_def) lemma rotate1_rl [simp] : "rotater1 (rotate1 l) = l" apply (unfold rotater1_def) apply (cases "l") apply (case_tac [2] "list") apply auto done lemma rotate1_lr [simp] : "rotate1 (rotater1 l) = l" by (cases l) (auto simp: rotater1_def) lemma rotater1_rev': "rotater1 (rev xs) = rev (rotate1 xs)" by (cases "xs") (simp add: rotater1_def, simp add: rotate1_rl') lemma rotater_rev': "rotater n (rev xs) = rev (rotate n xs)" by (induct n) (auto simp: rotater_def intro: rotater1_rev') lemma rotater_rev: "rotater n ys = rev (rotate n (rev ys))" using rotater_rev' [where xs = "rev ys"] by simp lemma rotater_drop_take: "rotater n xs = drop (length xs - n mod length xs) xs @ take (length xs - n mod length xs) xs" by (auto simp: rotater_rev rotate_drop_take rev_take rev_drop) lemma rotater_Suc [simp]: "rotater (Suc n) xs = rotater1 (rotater n xs)" unfolding rotater_def by auto lemma nth_rotater: \rotater m xs ! n = xs ! ((n + (length xs - m mod length xs)) mod length xs)\ if \n < length xs\ using that by (simp add: rotater_drop_take nth_append not_less less_diff_conv ac_simps le_mod_geq) lemma nth_rotater1: \rotater1 xs ! n = xs ! ((n + (length xs - 1)) mod length xs)\ if \n < length xs\ using that nth_rotater [of n xs 1] by simp lemma rotate_inv_plus [rule_format]: "\k. k = m + n \ rotater k (rotate n xs) = rotater m xs \ rotate k (rotater n xs) = rotate m xs \ rotater n (rotate k xs) = rotate m xs \ rotate n (rotater k xs) = rotater m xs" by (induct n) (auto simp: rotater_def rotate_def intro: funpow_swap1 [THEN trans]) lemmas rotate_inv_rel = le_add_diff_inverse2 [symmetric, THEN rotate_inv_plus] lemmas rotate_inv_eq = order_refl [THEN rotate_inv_rel, simplified] lemmas rotate_lr [simp] = rotate_inv_eq [THEN conjunct1] lemmas rotate_rl [simp] = rotate_inv_eq [THEN conjunct2, THEN conjunct1] lemma rotate_gal: "rotater n xs = ys \ rotate n ys = xs" by auto lemma rotate_gal': "ys = rotater n xs \ xs = rotate n ys" by auto lemma length_rotater [simp]: "length (rotater n xs) = length xs" by (simp add : rotater_rev) lemma rotate_eq_mod: "m mod length xs = n mod length xs \ rotate m xs = rotate n xs" apply (rule box_equals) defer apply (rule rotate_conv_mod [symmetric])+ apply simp done lemma restrict_to_left: "x = y \ x = z \ y = z" by simp lemmas rotate_eqs = trans [OF rotate0 [THEN fun_cong] id_apply] rotate_rotate [symmetric] rotate_id rotate_conv_mod rotate_eq_mod lemmas rrs0 = rotate_eqs [THEN restrict_to_left, simplified rotate_gal [symmetric] rotate_gal' [symmetric]] lemmas rrs1 = rrs0 [THEN refl [THEN rev_iffD1]] lemmas rotater_eqs = rrs1 [simplified length_rotater] lemmas rotater_0 = rotater_eqs (1) lemmas rotater_add = rotater_eqs (2) lemma butlast_map: "xs \ [] \ butlast (map f xs) = map f (butlast xs)" by (induct xs) auto lemma rotater1_map: "rotater1 (map f xs) = map f (rotater1 xs)" by (cases xs) (auto simp: rotater1_def last_map butlast_map) lemma rotater_map: "rotater n (map f xs) = map f (rotater n xs)" by (induct n) (auto simp: rotater_def rotater1_map) lemma but_last_zip [rule_format] : "\ys. length xs = length ys \ xs \ [] \ last (zip xs ys) = (last xs, last ys) \ butlast (zip xs ys) = zip (butlast xs) (butlast ys)" apply (induct xs) apply auto apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+ done lemma but_last_map2 [rule_format] : "\ys. length xs = length ys \ xs \ [] \ last (map2 f xs ys) = f (last xs) (last ys) \ butlast (map2 f xs ys) = map2 f (butlast xs) (butlast ys)" apply (induct xs) apply auto apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+ done lemma rotater1_zip: "length xs = length ys \ rotater1 (zip xs ys) = zip (rotater1 xs) (rotater1 ys)" apply (unfold rotater1_def) apply (cases xs) apply auto apply ((case_tac ys, auto simp: neq_Nil_conv but_last_zip)[1])+ done lemma rotater1_map2: "length xs = length ys \ rotater1 (map2 f xs ys) = map2 f (rotater1 xs) (rotater1 ys)" by (simp add: rotater1_map rotater1_zip) lemmas lrth = box_equals [OF asm_rl length_rotater [symmetric] length_rotater [symmetric], THEN rotater1_map2] lemma rotater_map2: "length xs = length ys \ rotater n (map2 f xs ys) = map2 f (rotater n xs) (rotater n ys)" by (induct n) (auto intro!: lrth) lemma rotate1_map2: "length xs = length ys \ rotate1 (map2 f xs ys) = map2 f (rotate1 xs) (rotate1 ys)" by (cases xs; cases ys) auto lemmas lth = box_equals [OF asm_rl length_rotate [symmetric] length_rotate [symmetric], THEN rotate1_map2] lemma rotate_map2: "length xs = length ys \ rotate n (map2 f xs ys) = map2 f (rotate n xs) (rotate n ys)" by (induct n) (auto intro!: lth) + +subsection \Explicit bit representation of \<^typ>\int\\ + +primrec bl_to_bin_aux :: "bool list \ int \ int" + where + Nil: "bl_to_bin_aux [] w = w" + | Cons: "bl_to_bin_aux (b # bs) w = bl_to_bin_aux bs (of_bool b + 2 * w)" + +definition bl_to_bin :: "bool list \ int" + where "bl_to_bin bs = bl_to_bin_aux bs 0" + +primrec bin_to_bl_aux :: "nat \ int \ bool list \ bool list" + where + Z: "bin_to_bl_aux 0 w bl = bl" + | Suc: "bin_to_bl_aux (Suc n) w bl = bin_to_bl_aux n (bin_rest w) ((bin_last w) # bl)" + +definition bin_to_bl :: "nat \ int \ bool list" + where "bin_to_bl n w = bin_to_bl_aux n w []" + +lemma bin_to_bl_aux_zero_minus_simp [simp]: + "0 < n \ bin_to_bl_aux n 0 bl = bin_to_bl_aux (n - 1) 0 (False # bl)" + by (cases n) auto + +lemma bin_to_bl_aux_minus1_minus_simp [simp]: + "0 < n \ bin_to_bl_aux n (- 1) bl = bin_to_bl_aux (n - 1) (- 1) (True # bl)" + by (cases n) auto + +lemma bin_to_bl_aux_one_minus_simp [simp]: + "0 < n \ bin_to_bl_aux n 1 bl = bin_to_bl_aux (n - 1) 0 (True # bl)" + by (cases n) auto + +lemma bin_to_bl_aux_Bit0_minus_simp [simp]: + "0 < n \ + bin_to_bl_aux n (numeral (Num.Bit0 w)) bl = bin_to_bl_aux (n - 1) (numeral w) (False # bl)" + by (cases n) simp_all + +lemma bin_to_bl_aux_Bit1_minus_simp [simp]: + "0 < n \ + bin_to_bl_aux n (numeral (Num.Bit1 w)) bl = bin_to_bl_aux (n - 1) (numeral w) (True # bl)" + by (cases n) simp_all + +lemma bl_to_bin_aux_append: "bl_to_bin_aux (bs @ cs) w = bl_to_bin_aux cs (bl_to_bin_aux bs w)" + by (induct bs arbitrary: w) auto + +lemma bin_to_bl_aux_append: "bin_to_bl_aux n w bs @ cs = bin_to_bl_aux n w (bs @ cs)" + by (induct n arbitrary: w bs) auto + +lemma bl_to_bin_append: "bl_to_bin (bs @ cs) = bl_to_bin_aux cs (bl_to_bin bs)" + unfolding bl_to_bin_def by (rule bl_to_bin_aux_append) + +lemma bin_to_bl_aux_alt: "bin_to_bl_aux n w bs = bin_to_bl n w @ bs" + by (simp add: bin_to_bl_def bin_to_bl_aux_append) + +lemma bin_to_bl_0 [simp]: "bin_to_bl 0 bs = []" + by (auto simp: bin_to_bl_def) + +lemma size_bin_to_bl_aux: "length (bin_to_bl_aux n w bs) = n + length bs" + by (induct n arbitrary: w bs) auto + +lemma size_bin_to_bl [simp]: "length (bin_to_bl n w) = n" + by (simp add: bin_to_bl_def size_bin_to_bl_aux) + +lemma bl_bin_bl': "bin_to_bl (n + length bs) (bl_to_bin_aux bs w) = bin_to_bl_aux n w bs" + apply (induct bs arbitrary: w n) + apply auto + apply (simp_all only: add_Suc [symmetric]) + apply (auto simp add: bin_to_bl_def) + done + +lemma bl_bin_bl [simp]: "bin_to_bl (length bs) (bl_to_bin bs) = bs" + unfolding bl_to_bin_def + apply (rule box_equals) + apply (rule bl_bin_bl') + prefer 2 + apply (rule bin_to_bl_aux.Z) + apply simp + done + +lemma bl_to_bin_inj: "bl_to_bin bs = bl_to_bin cs \ length bs = length cs \ bs = cs" + apply (rule_tac box_equals) + defer + apply (rule bl_bin_bl) + apply (rule bl_bin_bl) + apply simp + done + +lemma bl_to_bin_False [simp]: "bl_to_bin (False # bl) = bl_to_bin bl" + by (auto simp: bl_to_bin_def) + +lemma bl_to_bin_Nil [simp]: "bl_to_bin [] = 0" + by (auto simp: bl_to_bin_def) + +lemma bin_to_bl_zero_aux: "bin_to_bl_aux n 0 bl = replicate n False @ bl" + by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same) + +lemma bin_to_bl_zero: "bin_to_bl n 0 = replicate n False" + by (simp add: bin_to_bl_def bin_to_bl_zero_aux) + +lemma bin_to_bl_minus1_aux: "bin_to_bl_aux n (- 1) bl = replicate n True @ bl" + by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same) + +lemma bin_to_bl_minus1: "bin_to_bl n (- 1) = replicate n True" + by (simp add: bin_to_bl_def bin_to_bl_minus1_aux) + + +subsection \Semantic interpretation of \<^typ>\bool list\ as \<^typ>\int\\ + +lemma bin_bl_bin': "bl_to_bin (bin_to_bl_aux n w bs) = bl_to_bin_aux bs (bintrunc n w)" + by (induct n arbitrary: w bs) (auto simp: bl_to_bin_def take_bit_Suc ac_simps mod_2_eq_odd) + +lemma bin_bl_bin [simp]: "bl_to_bin (bin_to_bl n w) = bintrunc n w" + by (auto simp: bin_to_bl_def bin_bl_bin') + +lemma bl_to_bin_rep_F: "bl_to_bin (replicate n False @ bl) = bl_to_bin bl" + by (simp add: bin_to_bl_zero_aux [symmetric] bin_bl_bin') (simp add: bl_to_bin_def) + +lemma bin_to_bl_trunc [simp]: "n \ m \ bin_to_bl n (bintrunc m w) = bin_to_bl n w" + by (auto intro: bl_to_bin_inj) + +lemma bin_to_bl_aux_bintr: + "bin_to_bl_aux n (bintrunc m bin) bl = + replicate (n - m) False @ bin_to_bl_aux (min n m) bin bl" + apply (induct n arbitrary: m bin bl) + apply clarsimp + apply clarsimp + apply (case_tac "m") + apply (clarsimp simp: bin_to_bl_zero_aux) + apply (erule thin_rl) + apply (induct_tac n) + apply (auto simp add: take_bit_Suc) + done + +lemma bin_to_bl_bintr: + "bin_to_bl n (bintrunc m bin) = replicate (n - m) False @ bin_to_bl (min n m) bin" + unfolding bin_to_bl_def by (rule bin_to_bl_aux_bintr) + +lemma bl_to_bin_rep_False: "bl_to_bin (replicate n False) = 0" + by (induct n) auto + +lemma len_bin_to_bl_aux: "length (bin_to_bl_aux n w bs) = n + length bs" + by (fact size_bin_to_bl_aux) + +lemma len_bin_to_bl: "length (bin_to_bl n w) = n" + by (fact size_bin_to_bl) (* FIXME: duplicate *) + +lemma sign_bl_bin': "bin_sign (bl_to_bin_aux bs w) = bin_sign w" + by (induction bs arbitrary: w) (simp_all add: bin_sign_def) + +lemma sign_bl_bin: "bin_sign (bl_to_bin bs) = 0" + by (simp add: bl_to_bin_def sign_bl_bin') + +lemma bl_sbin_sign_aux: "hd (bin_to_bl_aux (Suc n) w bs) = (bin_sign (sbintrunc n w) = -1)" + by (induction n arbitrary: w bs) (auto simp add: bin_sign_def even_iff_mod_2_eq_zero bit_Suc) + +lemma bl_sbin_sign: "hd (bin_to_bl (Suc n) w) = (bin_sign (sbintrunc n w) = -1)" + unfolding bin_to_bl_def by (rule bl_sbin_sign_aux) + +lemma bin_nth_of_bl_aux: + "bin_nth (bl_to_bin_aux bl w) n = + (n < size bl \ rev bl ! n \ n \ length bl \ bin_nth w (n - size bl))" + apply (induction bl arbitrary: w) + apply simp_all + apply safe + apply (simp_all add: not_le nth_append bit_double_iff even_bit_succ_iff split: if_splits) + done + +lemma bin_nth_of_bl: "bin_nth (bl_to_bin bl) n = (n < length bl \ rev bl ! n)" + by (simp add: bl_to_bin_def bin_nth_of_bl_aux) + +lemma bin_nth_bl: "n < m \ bin_nth w n = nth (rev (bin_to_bl m w)) n" + apply (induct n arbitrary: m w) + apply clarsimp + apply (case_tac m, clarsimp) + apply (clarsimp simp: bin_to_bl_def) + apply (simp add: bin_to_bl_aux_alt) + apply (case_tac m, clarsimp) + apply (clarsimp simp: bin_to_bl_def) + apply (simp add: bin_to_bl_aux_alt bit_Suc) + done + +lemma nth_bin_to_bl_aux: + "n < m + length bl \ (bin_to_bl_aux m w bl) ! n = + (if n < m then bit w (m - 1 - n) else bl ! (n - m))" + apply (induction bl arbitrary: w) + apply simp_all + apply (simp add: bin_nth_bl [of \m - Suc n\ m] rev_nth flip: bin_to_bl_def) + apply (metis One_nat_def Suc_pred add_diff_cancel_left' + add_diff_cancel_right' bin_to_bl_aux_alt bin_to_bl_def + diff_Suc_Suc diff_is_0_eq diff_zero less_Suc_eq_0_disj + less_antisym less_imp_Suc_add list.size(3) nat_less_le nth_append size_bin_to_bl_aux) + done + +lemma nth_bin_to_bl: "n < m \ (bin_to_bl m w) ! n = bin_nth w (m - Suc n)" + by (simp add: bin_to_bl_def nth_bin_to_bl_aux) + +lemma takefill_bintrunc: "takefill False n bl = rev (bin_to_bl n (bl_to_bin (rev bl)))" + apply (rule nth_equalityI) + apply simp + apply (clarsimp simp: nth_takefill rev_nth nth_bin_to_bl bin_nth_of_bl) + done + +lemma bl_bin_bl_rtf: "bin_to_bl n (bl_to_bin bl) = rev (takefill False n (rev bl))" + by (simp add: takefill_bintrunc) + +lemma bl_to_bin_lt2p_aux: "bl_to_bin_aux bs w < (w + 1) * (2 ^ length bs)" +proof (induction bs arbitrary: w) + case Nil + then show ?case + by simp +next + case (Cons b bs) + from Cons.IH [of \1 + 2 * w\] Cons.IH [of \2 * w\] + show ?case + apply (auto simp add: algebra_simps) + apply (subst mult_2 [of \2 ^ length bs\]) + apply (simp only: add.assoc) + apply (rule pos_add_strict) + apply simp_all + done +qed + +lemma bl_to_bin_lt2p_drop: "bl_to_bin bs < 2 ^ length (dropWhile Not bs)" +proof (induct bs) + case Nil + then show ?case by simp +next + case (Cons b bs) + with bl_to_bin_lt2p_aux[where w=1] show ?case + by (simp add: bl_to_bin_def) +qed + +lemma bl_to_bin_lt2p: "bl_to_bin bs < 2 ^ length bs" + by (metis bin_bl_bin bintr_lt2p bl_bin_bl) + +lemma bl_to_bin_ge2p_aux: "bl_to_bin_aux bs w \ w * (2 ^ length bs)" +proof (induction bs arbitrary: w) + case Nil + then show ?case + by simp +next + case (Cons b bs) + from Cons.IH [of \1 + 2 * w\] Cons.IH [of \2 * w\] + show ?case + apply (auto simp add: algebra_simps) + apply (rule add_le_imp_le_left [of \2 ^ length bs\]) + apply (rule add_increasing) + apply simp_all + done +qed + +lemma bl_to_bin_ge0: "bl_to_bin bs \ 0" + apply (unfold bl_to_bin_def) + apply (rule xtrans(4)) + apply (rule bl_to_bin_ge2p_aux) + apply simp + done + +lemma butlast_rest_bin: "butlast (bin_to_bl n w) = bin_to_bl (n - 1) (bin_rest w)" + apply (unfold bin_to_bl_def) + apply (cases n, clarsimp) + apply clarsimp + apply (auto simp add: bin_to_bl_aux_alt) + done + +lemma butlast_bin_rest: "butlast bl = bin_to_bl (length bl - Suc 0) (bin_rest (bl_to_bin bl))" + using butlast_rest_bin [where w="bl_to_bin bl" and n="length bl"] by simp + +lemma butlast_rest_bl2bin_aux: + "bl \ [] \ bl_to_bin_aux (butlast bl) w = bin_rest (bl_to_bin_aux bl w)" + by (induct bl arbitrary: w) auto + +lemma butlast_rest_bl2bin: "bl_to_bin (butlast bl) = bin_rest (bl_to_bin bl)" + by (cases bl) (auto simp: bl_to_bin_def butlast_rest_bl2bin_aux) + +lemma trunc_bl2bin_aux: + "bintrunc m (bl_to_bin_aux bl w) = + bl_to_bin_aux (drop (length bl - m) bl) (bintrunc (m - length bl) w)" +proof (induct bl arbitrary: w) + case Nil + show ?case by simp +next + case (Cons b bl) + show ?case + proof (cases "m - length bl") + case 0 + then have "Suc (length bl) - m = Suc (length bl - m)" by simp + with Cons show ?thesis by simp + next + case (Suc n) + then have "m - Suc (length bl) = n" by simp + with Cons Suc show ?thesis by (simp add: take_bit_Suc ac_simps) + qed +qed + +lemma trunc_bl2bin: "bintrunc m (bl_to_bin bl) = bl_to_bin (drop (length bl - m) bl)" + by (simp add: bl_to_bin_def trunc_bl2bin_aux) + +lemma trunc_bl2bin_len [simp]: "bintrunc (length bl) (bl_to_bin bl) = bl_to_bin bl" + by (simp add: trunc_bl2bin) + +lemma bl2bin_drop: "bl_to_bin (drop k bl) = bintrunc (length bl - k) (bl_to_bin bl)" + apply (rule trans) + prefer 2 + apply (rule trunc_bl2bin [symmetric]) + apply (cases "k \ length bl") + apply auto + done + +lemma take_rest_power_bin: "m \ n \ take m (bin_to_bl n w) = bin_to_bl m ((bin_rest ^^ (n - m)) w)" + apply (rule nth_equalityI) + apply simp + apply (clarsimp simp add: nth_bin_to_bl nth_rest_power_bin) + done + +lemma last_bin_last': "size xs > 0 \ last xs \ bin_last (bl_to_bin_aux xs w)" + by (induct xs arbitrary: w) auto + +lemma last_bin_last: "size xs > 0 \ last xs \ bin_last (bl_to_bin xs)" + unfolding bl_to_bin_def by (erule last_bin_last') + +lemma bin_last_last: "bin_last w \ last (bin_to_bl (Suc n) w)" + by (simp add: bin_to_bl_def) (auto simp: bin_to_bl_aux_alt) + +lemma drop_bin2bl_aux: + "drop m (bin_to_bl_aux n bin bs) = + bin_to_bl_aux (n - m) bin (drop (m - n) bs)" + apply (induction n arbitrary: m bin bs) + apply auto + apply (case_tac "m \ n") + apply (auto simp add: not_le Suc_diff_le) + apply (case_tac "m - n") + apply auto + apply (use Suc_diff_Suc in fastforce) + done + +lemma drop_bin2bl: "drop m (bin_to_bl n bin) = bin_to_bl (n - m) bin" + by (simp add: bin_to_bl_def drop_bin2bl_aux) + +lemma take_bin2bl_lem1: "take m (bin_to_bl_aux m w bs) = bin_to_bl m w" + apply (induct m arbitrary: w bs) + apply clarsimp + apply clarsimp + apply (simp add: bin_to_bl_aux_alt) + apply (simp add: bin_to_bl_def) + apply (simp add: bin_to_bl_aux_alt) + done + +lemma take_bin2bl_lem: "take m (bin_to_bl_aux (m + n) w bs) = take m (bin_to_bl (m + n) w)" + by (induct n arbitrary: w bs) (simp_all (no_asm) add: bin_to_bl_def take_bin2bl_lem1, simp) + +lemma bin_split_take: "bin_split n c = (a, b) \ bin_to_bl m a = take m (bin_to_bl (m + n) c)" + apply (induct n arbitrary: b c) + apply clarsimp + apply (clarsimp simp: Let_def split: prod.split_asm) + apply (simp add: bin_to_bl_def) + apply (simp add: take_bin2bl_lem drop_bit_Suc) + done + +lemma bin_to_bl_drop_bit: + "k = m + n \ bin_to_bl m (drop_bit n c) = take m (bin_to_bl k c)" + using bin_split_take by simp + +lemma bin_split_take1: + "k = m + n \ bin_split n c = (a, b) \ bin_to_bl m a = take m (bin_to_bl k c)" + using bin_split_take by simp + +lemma bl_bin_bl_rep_drop: + "bin_to_bl n (bl_to_bin bl) = + replicate (n - length bl) False @ drop (length bl - n) bl" + by (simp add: bl_to_bin_inj bl_to_bin_rep_F trunc_bl2bin) + +lemma bl_to_bin_aux_cat: + "bl_to_bin_aux bs (bin_cat w nv v) = + bin_cat w (nv + length bs) (bl_to_bin_aux bs v)" + by (rule bit_eqI) + (auto simp add: bin_nth_of_bl_aux bin_nth_cat algebra_simps) + +lemma bin_to_bl_aux_cat: + "bin_to_bl_aux (nv + nw) (bin_cat v nw w) bs = + bin_to_bl_aux nv v (bin_to_bl_aux nw w bs)" + by (induction nw arbitrary: w bs) (simp_all add: concat_bit_Suc) + +lemma bl_to_bin_aux_alt: "bl_to_bin_aux bs w = bin_cat w (length bs) (bl_to_bin bs)" + using bl_to_bin_aux_cat [where nv = "0" and v = "0"] + by (simp add: bl_to_bin_def [symmetric]) + +lemma bin_to_bl_cat: + "bin_to_bl (nv + nw) (bin_cat v nw w) = + bin_to_bl_aux nv v (bin_to_bl nw w)" + by (simp add: bin_to_bl_def bin_to_bl_aux_cat) + +lemmas bl_to_bin_aux_app_cat = + trans [OF bl_to_bin_aux_append bl_to_bin_aux_alt] + +lemmas bin_to_bl_aux_cat_app = + trans [OF bin_to_bl_aux_cat bin_to_bl_aux_alt] + +lemma bl_to_bin_app_cat: + "bl_to_bin (bsa @ bs) = bin_cat (bl_to_bin bsa) (length bs) (bl_to_bin bs)" + by (simp only: bl_to_bin_aux_app_cat bl_to_bin_def) + +lemma bin_to_bl_cat_app: + "bin_to_bl (n + nw) (bin_cat w nw wa) = bin_to_bl n w @ bin_to_bl nw wa" + by (simp only: bin_to_bl_def bin_to_bl_aux_cat_app) + +text \\bl_to_bin_app_cat_alt\ and \bl_to_bin_app_cat\ are easily interderivable.\ +lemma bl_to_bin_app_cat_alt: "bin_cat (bl_to_bin cs) n w = bl_to_bin (cs @ bin_to_bl n w)" + by (simp add: bl_to_bin_app_cat) + +lemma mask_lem: "(bl_to_bin (True # replicate n False)) = bl_to_bin (replicate n True) + 1" + apply (unfold bl_to_bin_def) + apply (induct n) + apply simp + apply (simp only: Suc_eq_plus1 replicate_add append_Cons [symmetric] bl_to_bin_aux_append) + apply simp + done + +lemma bin_exhaust: + "(\x b. bin = of_bool b + 2 * x \ Q) \ Q" for bin :: int + apply (cases \even bin\) + apply (auto elim!: evenE oddE) + apply fastforce + apply fastforce + done + +primrec rbl_succ :: "bool list \ bool list" + where + Nil: "rbl_succ Nil = Nil" + | Cons: "rbl_succ (x # xs) = (if x then False # rbl_succ xs else True # xs)" + +primrec rbl_pred :: "bool list \ bool list" + where + Nil: "rbl_pred Nil = Nil" + | Cons: "rbl_pred (x # xs) = (if x then False # xs else True # rbl_pred xs)" + +primrec rbl_add :: "bool list \ bool list \ bool list" + where \ \result is length of first arg, second arg may be longer\ + Nil: "rbl_add Nil x = Nil" + | Cons: "rbl_add (y # ys) x = + (let ws = rbl_add ys (tl x) + in (y \ hd x) # (if hd x \ y then rbl_succ ws else ws))" + +primrec rbl_mult :: "bool list \ bool list \ bool list" + where \ \result is length of first arg, second arg may be longer\ + Nil: "rbl_mult Nil x = Nil" + | Cons: "rbl_mult (y # ys) x = + (let ws = False # rbl_mult ys x + in if y then rbl_add ws x else ws)" + +lemma size_rbl_pred: "length (rbl_pred bl) = length bl" + by (induct bl) auto + +lemma size_rbl_succ: "length (rbl_succ bl) = length bl" + by (induct bl) auto + +lemma size_rbl_add: "length (rbl_add bl cl) = length bl" + by (induct bl arbitrary: cl) (auto simp: Let_def size_rbl_succ) + +lemma size_rbl_mult: "length (rbl_mult bl cl) = length bl" + by (induct bl arbitrary: cl) (auto simp add: Let_def size_rbl_add) + +lemmas rbl_sizes [simp] = + size_rbl_pred size_rbl_succ size_rbl_add size_rbl_mult + +lemmas rbl_Nils = + rbl_pred.Nil rbl_succ.Nil rbl_add.Nil rbl_mult.Nil + +lemma rbl_add_app2: "length blb \ length bla \ rbl_add bla (blb @ blc) = rbl_add bla blb" + apply (induct bla arbitrary: blb) + apply simp + apply clarsimp + apply (case_tac blb, clarsimp) + apply (clarsimp simp: Let_def) + done + +lemma rbl_add_take2: + "length blb \ length bla \ rbl_add bla (take (length bla) blb) = rbl_add bla blb" + apply (induct bla arbitrary: blb) + apply simp + apply clarsimp + apply (case_tac blb, clarsimp) + apply (clarsimp simp: Let_def) + done + +lemma rbl_mult_app2: "length blb \ length bla \ rbl_mult bla (blb @ blc) = rbl_mult bla blb" + apply (induct bla arbitrary: blb) + apply simp + apply clarsimp + apply (case_tac blb, clarsimp) + apply (clarsimp simp: Let_def rbl_add_app2) + done + +lemma rbl_mult_take2: + "length blb \ length bla \ rbl_mult bla (take (length bla) blb) = rbl_mult bla blb" + apply (rule trans) + apply (rule rbl_mult_app2 [symmetric]) + apply simp + apply (rule_tac f = "rbl_mult bla" in arg_cong) + apply (rule append_take_drop_id) + done + +lemma rbl_add_split: + "P (rbl_add (y # ys) (x # xs)) = + (\ws. length ws = length ys \ ws = rbl_add ys xs \ + (y \ ((x \ P (False # rbl_succ ws)) \ (\ x \ P (True # ws)))) \ + (\ y \ P (x # ws)))" + by (cases y) (auto simp: Let_def) + +lemma rbl_mult_split: + "P (rbl_mult (y # ys) xs) = + (\ws. length ws = Suc (length ys) \ ws = False # rbl_mult ys xs \ + (y \ P (rbl_add ws xs)) \ (\ y \ P ws))" + by (auto simp: Let_def) + +lemma rbl_pred: "rbl_pred (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin - 1))" +proof (unfold bin_to_bl_def, induction n arbitrary: bin) + case 0 + then show ?case + by simp +next + case (Suc n) + obtain b k where \bin = of_bool b + 2 * k\ + using bin_exhaust by blast + moreover have \(2 * k - 1) div 2 = k - 1\ + using even_succ_div_2 [of \2 * (k - 1)\] + by simp + ultimately show ?case + using Suc [of \bin div 2\] + by simp (simp add: bin_to_bl_aux_alt) +qed + +lemma rbl_succ: "rbl_succ (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin + 1))" + apply (unfold bin_to_bl_def) + apply (induction n arbitrary: bin) + apply simp_all + apply (case_tac bin rule: bin_exhaust) + apply simp + apply (simp add: bin_to_bl_aux_alt ac_simps) + done + +lemma rbl_add: + "\bina binb. rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = + rev (bin_to_bl n (bina + binb))" + apply (unfold bin_to_bl_def) + apply (induct n) + apply simp + apply clarsimp + apply (case_tac bina rule: bin_exhaust) + apply (case_tac binb rule: bin_exhaust) + apply (case_tac b) + apply (case_tac [!] "ba") + apply (auto simp: rbl_succ bin_to_bl_aux_alt Let_def ac_simps) + done + +lemma rbl_add_long: + "m \ n \ rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = + rev (bin_to_bl n (bina + binb))" + apply (rule box_equals [OF _ rbl_add_take2 rbl_add]) + apply (rule_tac f = "rbl_add (rev (bin_to_bl n bina))" in arg_cong) + apply (rule rev_swap [THEN iffD1]) + apply (simp add: rev_take drop_bin2bl) + apply simp + done + +lemma rbl_mult_gt1: + "m \ length bl \ + rbl_mult bl (rev (bin_to_bl m binb)) = + rbl_mult bl (rev (bin_to_bl (length bl) binb))" + apply (rule trans) + apply (rule rbl_mult_take2 [symmetric]) + apply simp_all + apply (rule_tac f = "rbl_mult bl" in arg_cong) + apply (rule rev_swap [THEN iffD1]) + apply (simp add: rev_take drop_bin2bl) + done + +lemma rbl_mult_gt: + "m > n \ + rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = + rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb))" + by (auto intro: trans [OF rbl_mult_gt1]) + +lemmas rbl_mult_Suc = lessI [THEN rbl_mult_gt] + +lemma rbbl_Cons: "b # rev (bin_to_bl n x) = rev (bin_to_bl (Suc n) (of_bool b + 2 * x))" + by (simp add: bin_to_bl_def) (simp add: bin_to_bl_aux_alt) + +lemma rbl_mult: + "rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = + rev (bin_to_bl n (bina * binb))" + apply (induct n arbitrary: bina binb) + apply simp_all + apply (unfold bin_to_bl_def) + apply clarsimp + apply (case_tac bina rule: bin_exhaust) + apply (case_tac binb rule: bin_exhaust) + apply simp + apply (simp add: bin_to_bl_aux_alt) + apply (simp add: rbbl_Cons rbl_mult_Suc rbl_add algebra_simps) + done + +lemma sclem: "size (concat (map (bin_to_bl n) xs)) = length xs * n" + by (induct xs) auto + +lemma bin_cat_foldl_lem: + "foldl (\u. bin_cat u n) x xs = + bin_cat x (size xs * n) (foldl (\u. bin_cat u n) y xs)" + apply (induct xs arbitrary: x) + apply simp + apply (simp (no_asm)) + apply (frule asm_rl) + apply (drule meta_spec) + apply (erule trans) + apply (drule_tac x = "bin_cat y n a" in meta_spec) + apply (simp add: bin_cat_assoc_sym min.absorb2) + done + +lemma bin_rcat_bl: "bin_rcat n wl = bl_to_bin (concat (map (bin_to_bl n) wl))" + apply (unfold bin_rcat_def) + apply (rule sym) + apply (induct wl) + apply (auto simp add: bl_to_bin_append) + apply (simp add: bl_to_bin_aux_alt sclem) + apply (simp add: bin_cat_foldl_lem [symmetric]) + done + +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 bl_xor_aux_bin: + "map2 (\x y. x \ y) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = + bin_to_bl_aux n (v XOR w) (map2 (\x y. x \ y) bs cs)" + apply (induction n arbitrary: v w bs cs) + apply auto + apply (case_tac v rule: bin_exhaust) + apply (case_tac w rule: bin_exhaust) + apply clarsimp + done + +lemma bl_or_aux_bin: + "map2 (\) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = + bin_to_bl_aux n (v OR w) (map2 (\) bs cs)" + by (induct n arbitrary: v w bs cs) simp_all + +lemma bl_and_aux_bin: + "map2 (\) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = + bin_to_bl_aux n (v AND w) (map2 (\) bs cs)" + by (induction n arbitrary: v w bs cs) simp_all + +lemma bl_not_aux_bin: "map Not (bin_to_bl_aux n w cs) = bin_to_bl_aux n (NOT w) (map Not cs)" + by (induct n arbitrary: w cs) auto + +lemma bl_not_bin: "map Not (bin_to_bl n w) = bin_to_bl n (NOT w)" + by (simp add: bin_to_bl_def bl_not_aux_bin) + +lemma bl_and_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v AND w)" + by (simp add: bin_to_bl_def bl_and_aux_bin) + +lemma bl_or_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v OR w)" + by (simp add: bin_to_bl_def bl_or_aux_bin) + +lemma bl_xor_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v XOR w)" + using bl_xor_aux_bin by (simp add: bin_to_bl_def) + end diff --git a/src/HOL/Word/Bits_Int.thy b/src/HOL/Word/Bits_Int.thy --- a/src/HOL/Word/Bits_Int.thy +++ b/src/HOL/Word/Bits_Int.thy @@ -1,2197 +1,1539 @@ (* Title: HOL/Word/Bits_Int.thy Author: Jeremy Dawson and Gerwin Klein, NICTA *) section \Bitwise Operations on integers\ theory Bits_Int imports "HOL-Library.Bit_Operations" Traditional_Syntax begin subsection \Generic auxiliary\ lemma int_mod_ge: "a < n \ 0 < n \ a \ a mod n" for a n :: int by (metis dual_order.trans le_cases mod_pos_pos_trivial pos_mod_conj) subsection \Implicit bit representation of \<^typ>\int\\ abbreviation (input) bin_last :: "int \ bool" where "bin_last \ odd" lemma bin_last_def: "bin_last w \ w mod 2 = 1" by (fact odd_iff_mod_2_eq_one) abbreviation (input) bin_rest :: "int \ int" where "bin_rest w \ w div 2" lemma bin_last_numeral_simps [simp]: "\ bin_last 0" "bin_last 1" "bin_last (- 1)" "bin_last Numeral1" "\ bin_last (numeral (Num.Bit0 w))" "bin_last (numeral (Num.Bit1 w))" "\ bin_last (- numeral (Num.Bit0 w))" "bin_last (- numeral (Num.Bit1 w))" by simp_all lemma bin_rest_numeral_simps [simp]: "bin_rest 0 = 0" "bin_rest 1 = 0" "bin_rest (- 1) = - 1" "bin_rest Numeral1 = 0" "bin_rest (numeral (Num.Bit0 w)) = numeral w" "bin_rest (numeral (Num.Bit1 w)) = numeral w" "bin_rest (- numeral (Num.Bit0 w)) = - numeral w" "bin_rest (- numeral (Num.Bit1 w)) = - numeral (w + Num.One)" by simp_all lemma bin_rl_eqI: "\bin_rest x = bin_rest y; bin_last x = bin_last y\ \ x = y" by (auto elim: oddE) 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 lemma bin_rest_gt_0 [simp]: "bin_rest x > 0 \ x > 1" by auto -subsection \Explicit bit representation of \<^typ>\int\\ - -primrec bl_to_bin_aux :: "bool list \ int \ int" - where - Nil: "bl_to_bin_aux [] w = w" - | Cons: "bl_to_bin_aux (b # bs) w = bl_to_bin_aux bs (of_bool b + 2 * w)" - -definition bl_to_bin :: "bool list \ int" - where "bl_to_bin bs = bl_to_bin_aux bs 0" - -primrec bin_to_bl_aux :: "nat \ int \ bool list \ bool list" - where - Z: "bin_to_bl_aux 0 w bl = bl" - | Suc: "bin_to_bl_aux (Suc n) w bl = bin_to_bl_aux n (bin_rest w) ((bin_last w) # bl)" - -definition bin_to_bl :: "nat \ int \ bool list" - where "bin_to_bl n w = bin_to_bl_aux n w []" - -lemma bin_to_bl_aux_zero_minus_simp [simp]: - "0 < n \ bin_to_bl_aux n 0 bl = bin_to_bl_aux (n - 1) 0 (False # bl)" - by (cases n) auto - -lemma bin_to_bl_aux_minus1_minus_simp [simp]: - "0 < n \ bin_to_bl_aux n (- 1) bl = bin_to_bl_aux (n - 1) (- 1) (True # bl)" - by (cases n) auto - -lemma bin_to_bl_aux_one_minus_simp [simp]: - "0 < n \ bin_to_bl_aux n 1 bl = bin_to_bl_aux (n - 1) 0 (True # bl)" - by (cases n) auto - -lemma bin_to_bl_aux_Bit0_minus_simp [simp]: - "0 < n \ - bin_to_bl_aux n (numeral (Num.Bit0 w)) bl = bin_to_bl_aux (n - 1) (numeral w) (False # bl)" - by (cases n) simp_all - -lemma bin_to_bl_aux_Bit1_minus_simp [simp]: - "0 < n \ - bin_to_bl_aux n (numeral (Num.Bit1 w)) bl = bin_to_bl_aux (n - 1) (numeral w) (True # bl)" - by (cases n) simp_all - -lemma bl_to_bin_aux_append: "bl_to_bin_aux (bs @ cs) w = bl_to_bin_aux cs (bl_to_bin_aux bs w)" - by (induct bs arbitrary: w) auto - -lemma bin_to_bl_aux_append: "bin_to_bl_aux n w bs @ cs = bin_to_bl_aux n w (bs @ cs)" - by (induct n arbitrary: w bs) auto - -lemma bl_to_bin_append: "bl_to_bin (bs @ cs) = bl_to_bin_aux cs (bl_to_bin bs)" - unfolding bl_to_bin_def by (rule bl_to_bin_aux_append) - -lemma bin_to_bl_aux_alt: "bin_to_bl_aux n w bs = bin_to_bl n w @ bs" - by (simp add: bin_to_bl_def bin_to_bl_aux_append) - -lemma bin_to_bl_0 [simp]: "bin_to_bl 0 bs = []" - by (auto simp: bin_to_bl_def) - -lemma size_bin_to_bl_aux: "length (bin_to_bl_aux n w bs) = n + length bs" - by (induct n arbitrary: w bs) auto - -lemma size_bin_to_bl [simp]: "length (bin_to_bl n w) = n" - by (simp add: bin_to_bl_def size_bin_to_bl_aux) - -lemma bl_bin_bl': "bin_to_bl (n + length bs) (bl_to_bin_aux bs w) = bin_to_bl_aux n w bs" - apply (induct bs arbitrary: w n) - apply auto - apply (simp_all only: add_Suc [symmetric]) - apply (auto simp add: bin_to_bl_def) - done - -lemma bl_bin_bl [simp]: "bin_to_bl (length bs) (bl_to_bin bs) = bs" - unfolding bl_to_bin_def - apply (rule box_equals) - apply (rule bl_bin_bl') - prefer 2 - apply (rule bin_to_bl_aux.Z) - apply simp - done - -lemma bl_to_bin_inj: "bl_to_bin bs = bl_to_bin cs \ length bs = length cs \ bs = cs" - apply (rule_tac box_equals) - defer - apply (rule bl_bin_bl) - apply (rule bl_bin_bl) - apply simp - done - -lemma bl_to_bin_False [simp]: "bl_to_bin (False # bl) = bl_to_bin bl" - by (auto simp: bl_to_bin_def) - -lemma bl_to_bin_Nil [simp]: "bl_to_bin [] = 0" - by (auto simp: bl_to_bin_def) - -lemma bin_to_bl_zero_aux: "bin_to_bl_aux n 0 bl = replicate n False @ bl" - by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same) - -lemma bin_to_bl_zero: "bin_to_bl n 0 = replicate n False" - by (simp add: bin_to_bl_def bin_to_bl_zero_aux) - -lemma bin_to_bl_minus1_aux: "bin_to_bl_aux n (- 1) bl = replicate n True @ bl" - by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same) - -lemma bin_to_bl_minus1: "bin_to_bl n (- 1) = replicate n True" - by (simp add: bin_to_bl_def bin_to_bl_minus1_aux) - - subsection \Bit projection\ abbreviation (input) bin_nth :: \int \ nat \ bool\ where \bin_nth \ bit\ lemma bin_nth_eq_iff: "bin_nth x = bin_nth y \ x = y" by (simp add: bit_eq_iff fun_eq_iff) lemma bin_eqI: "x = y" if "\n. bin_nth x n \ bin_nth y n" using that bin_nth_eq_iff [of x y] by (simp add: fun_eq_iff) lemma bin_eq_iff: "x = y \ (\n. bin_nth x n = bin_nth y n)" by (fact bit_eq_iff) lemma bin_nth_zero [simp]: "\ bin_nth 0 n" by simp lemma bin_nth_1 [simp]: "bin_nth 1 n \ n = 0" by (cases n) (simp_all add: bit_Suc) lemma bin_nth_minus1 [simp]: "bin_nth (- 1) n" by (induction n) (simp_all add: bit_Suc) lemma bin_nth_numeral: "bin_rest x = y \ bin_nth x (numeral n) = bin_nth y (pred_numeral n)" by (simp add: numeral_eq_Suc bit_Suc) lemmas bin_nth_numeral_simps [simp] = bin_nth_numeral [OF bin_rest_numeral_simps(2)] bin_nth_numeral [OF bin_rest_numeral_simps(5)] bin_nth_numeral [OF bin_rest_numeral_simps(6)] bin_nth_numeral [OF bin_rest_numeral_simps(7)] bin_nth_numeral [OF bin_rest_numeral_simps(8)] lemmas bin_nth_simps = bit_0 bit_Suc bin_nth_zero bin_nth_minus1 bin_nth_numeral_simps lemma nth_2p_bin: "bin_nth (2 ^ n) m = (m = n)" \ \for use when simplifying with \bin_nth_Bit\\ by (auto simp add: bit_exp_iff) lemma nth_rest_power_bin: "bin_nth ((bin_rest ^^ k) w) n = bin_nth w (n + k)" apply (induct k arbitrary: n) apply clarsimp apply clarsimp apply (simp only: bit_Suc [symmetric] add_Suc) done 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 (cases n; simp)+ subsection \Truncating\ definition bin_sign :: "int \ int" where "bin_sign k = (if k \ 0 then 0 else - 1)" lemma bin_sign_simps [simp]: "bin_sign 0 = 0" "bin_sign 1 = 0" "bin_sign (- 1) = - 1" "bin_sign (numeral k) = 0" "bin_sign (- numeral k) = -1" by (simp_all add: bin_sign_def) lemma bin_sign_rest [simp]: "bin_sign (bin_rest w) = bin_sign w" by (simp add: bin_sign_def) abbreviation (input) bintrunc :: \nat \ int \ int\ where \bintrunc \ take_bit\ lemma bintrunc_mod2p: "bintrunc n w = w mod 2 ^ n" by (fact take_bit_eq_mod) abbreviation (input) sbintrunc :: \nat \ int \ int\ where \sbintrunc \ signed_take_bit\ abbreviation (input) norm_sint :: \nat \ int \ int\ where \norm_sint n \ signed_take_bit (n - 1)\ lemma sbintrunc_mod2p: "sbintrunc n w = (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n" by (simp add: bintrunc_mod2p signed_take_bit_eq_take_bit_shift) lemma sbintrunc_eq_take_bit: \sbintrunc n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n\ by (fact signed_take_bit_eq_take_bit_shift) lemma sign_bintr: "bin_sign (bintrunc n w) = 0" by (simp add: bin_sign_def) lemma bintrunc_n_0: "bintrunc n 0 = 0" by (fact take_bit_of_0) lemma sbintrunc_n_0: "sbintrunc n 0 = 0" by (fact signed_take_bit_of_0) lemma sbintrunc_n_minus1: "sbintrunc n (- 1) = -1" by (fact signed_take_bit_of_minus_1) lemma bintrunc_Suc_numeral: "bintrunc (Suc n) 1 = 1" "bintrunc (Suc n) (- 1) = 1 + 2 * bintrunc n (- 1)" "bintrunc (Suc n) (numeral (Num.Bit0 w)) = 2 * bintrunc n (numeral w)" "bintrunc (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * bintrunc n (numeral w)" "bintrunc (Suc n) (- numeral (Num.Bit0 w)) = 2 * bintrunc n (- numeral w)" "bintrunc (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * bintrunc n (- numeral (w + Num.One))" by (simp_all add: take_bit_Suc) lemma sbintrunc_0_numeral [simp]: "sbintrunc 0 1 = -1" "sbintrunc 0 (numeral (Num.Bit0 w)) = 0" "sbintrunc 0 (numeral (Num.Bit1 w)) = -1" "sbintrunc 0 (- numeral (Num.Bit0 w)) = 0" "sbintrunc 0 (- numeral (Num.Bit1 w)) = -1" by simp_all lemma sbintrunc_Suc_numeral: "sbintrunc (Suc n) 1 = 1" "sbintrunc (Suc n) (numeral (Num.Bit0 w)) = 2 * sbintrunc n (numeral w)" "sbintrunc (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * sbintrunc n (numeral w)" "sbintrunc (Suc n) (- numeral (Num.Bit0 w)) = 2 * sbintrunc n (- numeral w)" "sbintrunc (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * sbintrunc n (- numeral (w + Num.One))" by (simp_all add: signed_take_bit_Suc) lemma bin_sign_lem: "(bin_sign (sbintrunc n bin) = -1) = bit bin n" by (simp add: bin_sign_def) lemma nth_bintr: "bin_nth (bintrunc m w) n \ n < m \ bin_nth w n" by (fact bit_take_bit_iff) lemma nth_sbintr: "bin_nth (sbintrunc m w) n = (if n < m then bin_nth w n else bin_nth w m)" by (simp add: bit_signed_take_bit_iff min_def) lemma bin_nth_Bit0: "bin_nth (numeral (Num.Bit0 w)) n \ (\m. n = Suc m \ bin_nth (numeral w) m)" using bit_double_iff [of \numeral w :: int\ n] by (auto intro: exI [of _ \n - 1\]) lemma bin_nth_Bit1: "bin_nth (numeral (Num.Bit1 w)) n \ n = 0 \ (\m. n = Suc m \ bin_nth (numeral w) m)" using even_bit_succ_iff [of \2 * numeral w :: int\ n] bit_double_iff [of \numeral w :: int\ n] by auto lemma bintrunc_bintrunc_l: "n \ m \ bintrunc m (bintrunc n w) = bintrunc n w" by (simp add: min.absorb2) lemma sbintrunc_sbintrunc_l: "n \ m \ sbintrunc m (sbintrunc n w) = sbintrunc n w" by (simp add: min_def) lemma bintrunc_bintrunc_ge: "n \ m \ bintrunc n (bintrunc m w) = bintrunc n w" by (rule bin_eqI) (auto simp: nth_bintr) lemma bintrunc_bintrunc_min [simp]: "bintrunc m (bintrunc n w) = bintrunc (min m n) w" by (rule bin_eqI) (auto simp: nth_bintr) lemma sbintrunc_sbintrunc_min [simp]: "sbintrunc m (sbintrunc n w) = sbintrunc (min m n) w" by (rule bin_eqI) (auto simp: nth_sbintr min.absorb1 min.absorb2) lemmas sbintrunc_Suc_Pls = signed_take_bit_Suc [where k="0", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Suc_Min = signed_take_bit_Suc [where k="-1", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Sucs = sbintrunc_Suc_Pls sbintrunc_Suc_Min sbintrunc_Suc_numeral lemmas sbintrunc_Pls = signed_take_bit_0 [where k="0", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Min = signed_take_bit_0 [where k="-1", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_0_simps = sbintrunc_Pls sbintrunc_Min lemmas sbintrunc_simps = sbintrunc_0_simps sbintrunc_Sucs lemma bintrunc_minus: "0 < n \ bintrunc (Suc (n - 1)) w = bintrunc n w" by auto lemma sbintrunc_minus: "0 < n \ sbintrunc (Suc (n - 1)) w = sbintrunc n w" by auto lemmas sbintrunc_minus_simps = sbintrunc_Sucs [THEN [2] sbintrunc_minus [symmetric, THEN trans]] lemma sbintrunc_BIT_I: \0 < n \ sbintrunc (n - 1) 0 = y \ sbintrunc n 0 = 2 * y\ by simp lemma sbintrunc_Suc_Is: \sbintrunc n (- 1) = y \ sbintrunc (Suc n) (- 1) = 1 + 2 * y\ by auto lemma sbintrunc_Suc_lem: "sbintrunc (Suc n) x = y \ m = Suc n \ sbintrunc m x = y" by auto lemmas sbintrunc_Suc_Ialts = sbintrunc_Suc_Is [THEN sbintrunc_Suc_lem] lemma sbintrunc_bintrunc_lt: "m > n \ sbintrunc n (bintrunc m w) = sbintrunc n w" by (rule bin_eqI) (auto simp: nth_sbintr nth_bintr) lemma bintrunc_sbintrunc_le: "m \ Suc n \ bintrunc m (sbintrunc n w) = bintrunc m w" apply (rule bin_eqI) using le_Suc_eq less_Suc_eq_le apply (auto simp: nth_sbintr nth_bintr) done lemmas bintrunc_sbintrunc [simp] = order_refl [THEN bintrunc_sbintrunc_le] lemmas sbintrunc_bintrunc [simp] = lessI [THEN sbintrunc_bintrunc_lt] lemmas bintrunc_bintrunc [simp] = order_refl [THEN bintrunc_bintrunc_l] lemmas sbintrunc_sbintrunc [simp] = order_refl [THEN sbintrunc_sbintrunc_l] lemma bintrunc_sbintrunc' [simp]: "0 < n \ bintrunc n (sbintrunc (n - 1) w) = bintrunc n w" by (cases n) simp_all lemma sbintrunc_bintrunc' [simp]: "0 < n \ sbintrunc (n - 1) (bintrunc n w) = sbintrunc (n - 1) w" by (cases n) simp_all lemma bin_sbin_eq_iff: "bintrunc (Suc n) x = bintrunc (Suc n) y \ sbintrunc n x = sbintrunc n y" apply (rule iffI) apply (rule box_equals [OF _ sbintrunc_bintrunc sbintrunc_bintrunc]) apply simp apply (rule box_equals [OF _ bintrunc_sbintrunc bintrunc_sbintrunc]) apply simp done lemma bin_sbin_eq_iff': "0 < n \ bintrunc n x = bintrunc n y \ sbintrunc (n - 1) x = sbintrunc (n - 1) y" by (cases n) (simp_all add: bin_sbin_eq_iff) lemmas bintrunc_sbintruncS0 [simp] = bintrunc_sbintrunc' [unfolded One_nat_def] lemmas sbintrunc_bintruncS0 [simp] = sbintrunc_bintrunc' [unfolded One_nat_def] lemmas bintrunc_bintrunc_l' = le_add1 [THEN bintrunc_bintrunc_l] lemmas sbintrunc_sbintrunc_l' = le_add1 [THEN sbintrunc_sbintrunc_l] (* although bintrunc_minus_simps, if added to default simpset, tends to get applied where it's not wanted in developing the theories, we get a version for when the word length is given literally *) lemmas nat_non0_gr = trans [OF iszero_def [THEN Not_eq_iff [THEN iffD2]] refl] lemma bintrunc_numeral: "bintrunc (numeral k) x = of_bool (odd x) + 2 * bintrunc (pred_numeral k) (x div 2)" by (simp add: numeral_eq_Suc take_bit_Suc mod_2_eq_odd) lemma sbintrunc_numeral: "sbintrunc (numeral k) x = of_bool (odd x) + 2 * sbintrunc (pred_numeral k) (x div 2)" by (simp add: numeral_eq_Suc signed_take_bit_Suc mod2_eq_if) lemma bintrunc_numeral_simps [simp]: "bintrunc (numeral k) (numeral (Num.Bit0 w)) = 2 * bintrunc (pred_numeral k) (numeral w)" "bintrunc (numeral k) (numeral (Num.Bit1 w)) = 1 + 2 * bintrunc (pred_numeral k) (numeral w)" "bintrunc (numeral k) (- numeral (Num.Bit0 w)) = 2 * bintrunc (pred_numeral k) (- numeral w)" "bintrunc (numeral k) (- numeral (Num.Bit1 w)) = 1 + 2 * bintrunc (pred_numeral k) (- numeral (w + Num.One))" "bintrunc (numeral k) 1 = 1" by (simp_all add: bintrunc_numeral) lemma sbintrunc_numeral_simps [simp]: "sbintrunc (numeral k) (numeral (Num.Bit0 w)) = 2 * sbintrunc (pred_numeral k) (numeral w)" "sbintrunc (numeral k) (numeral (Num.Bit1 w)) = 1 + 2 * sbintrunc (pred_numeral k) (numeral w)" "sbintrunc (numeral k) (- numeral (Num.Bit0 w)) = 2 * sbintrunc (pred_numeral k) (- numeral w)" "sbintrunc (numeral k) (- numeral (Num.Bit1 w)) = 1 + 2 * sbintrunc (pred_numeral k) (- numeral (w + Num.One))" "sbintrunc (numeral k) 1 = 1" by (simp_all add: sbintrunc_numeral) lemma no_bintr_alt1: "bintrunc n = (\w. w mod 2 ^ n :: int)" by (rule ext) (rule bintrunc_mod2p) lemma range_bintrunc: "range (bintrunc n) = {i. 0 \ i \ i < 2 ^ n}" by (auto simp add: take_bit_eq_mod image_iff) (metis mod_pos_pos_trivial) lemma no_sbintr_alt2: "sbintrunc n = (\w. (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n :: int)" by (rule ext) (simp add : sbintrunc_mod2p) lemma range_sbintrunc: "range (sbintrunc n) = {i. - (2 ^ n) \ i \ i < 2 ^ n}" proof - have \surj (\k::int. k + 2 ^ n)\ by (rule surjI [of _ \(\k. k - 2 ^ n)\]) simp moreover have \sbintrunc n = ((\k. k - 2 ^ n) \ take_bit (Suc n) \ (\k. k + 2 ^ n))\ by (simp add: sbintrunc_eq_take_bit fun_eq_iff) ultimately show ?thesis apply (simp only: fun.set_map range_bintrunc) apply (auto simp add: image_iff) apply presburger done qed lemma sbintrunc_inc: \k + 2 ^ Suc n \ sbintrunc n k\ if \k < - (2 ^ n)\ using that by (fact signed_take_bit_greater_eq) lemma sbintrunc_dec: \sbintrunc n k \ k - 2 ^ (Suc n)\ if \k \ 2 ^ n\ using that by (fact signed_take_bit_less_eq) lemma bintr_ge0: "0 \ bintrunc n w" by (simp add: bintrunc_mod2p) lemma bintr_lt2p: "bintrunc n w < 2 ^ n" by (simp add: bintrunc_mod2p) lemma bintr_Min: "bintrunc n (- 1) = 2 ^ n - 1" by (simp add: stable_imp_take_bit_eq) lemma sbintr_ge: "- (2 ^ n) \ sbintrunc n w" by (simp add: sbintrunc_mod2p) lemma sbintr_lt: "sbintrunc n w < 2 ^ n" by (simp add: sbintrunc_mod2p) lemma sign_Pls_ge_0: "bin_sign bin = 0 \ bin \ 0" for bin :: int by (simp add: bin_sign_def) lemma sign_Min_lt_0: "bin_sign bin = -1 \ bin < 0" for bin :: int by (simp add: bin_sign_def) lemma bin_rest_trunc: "bin_rest (bintrunc n bin) = bintrunc (n - 1) (bin_rest bin)" by (simp add: take_bit_rec [of n bin]) lemma bin_rest_power_trunc: "(bin_rest ^^ k) (bintrunc n bin) = bintrunc (n - k) ((bin_rest ^^ k) bin)" by (induct k) (auto simp: bin_rest_trunc) lemma bin_rest_trunc_i: "bintrunc n (bin_rest bin) = bin_rest (bintrunc (Suc n) bin)" by (auto simp add: take_bit_Suc) lemma bin_rest_strunc: "bin_rest (sbintrunc (Suc n) bin) = sbintrunc n (bin_rest bin)" by (simp add: signed_take_bit_Suc) lemma bintrunc_rest [simp]: "bintrunc n (bin_rest (bintrunc n bin)) = bin_rest (bintrunc n bin)" by (induct n arbitrary: bin) (simp_all add: take_bit_Suc) lemma sbintrunc_rest [simp]: "sbintrunc n (bin_rest (sbintrunc n bin)) = bin_rest (sbintrunc n bin)" by (induct n arbitrary: bin) (simp_all add: signed_take_bit_Suc mod2_eq_if) lemma bintrunc_rest': "bintrunc n \ bin_rest \ bintrunc n = bin_rest \ bintrunc n" by (rule ext) auto lemma sbintrunc_rest': "sbintrunc n \ bin_rest \ sbintrunc n = bin_rest \ sbintrunc n" by (rule ext) auto lemma rco_lem: "f \ g \ f = g \ f \ f \ (g \ f) ^^ n = g ^^ n \ f" apply (rule ext) apply (induct_tac n) apply (simp_all (no_asm)) apply (drule fun_cong) apply (unfold o_def) apply (erule trans) apply simp done lemmas rco_bintr = bintrunc_rest' [THEN rco_lem [THEN fun_cong], unfolded o_def] lemmas rco_sbintr = sbintrunc_rest' [THEN rco_lem [THEN fun_cong], unfolded o_def] subsection \Splitting and concatenation\ definition bin_split :: \nat \ int \ int \ int\ where [simp]: \bin_split n k = (drop_bit n k, take_bit n k)\ lemma [code]: "bin_split (Suc n) w = (let (w1, w2) = bin_split n (w div 2) in (w1, of_bool (odd w) + 2 * w2))" "bin_split 0 w = (w, 0)" by (simp_all add: drop_bit_Suc take_bit_Suc mod_2_eq_odd) abbreviation (input) bin_cat :: \int \ nat \ int \ int\ where \bin_cat k n l \ concat_bit n l k\ lemma bin_cat_eq_push_bit_add_take_bit: \bin_cat k n l = push_bit n k + take_bit n l\ by (simp add: concat_bit_eq) lemma bin_sign_cat: "bin_sign (bin_cat x n y) = bin_sign x" proof - have \0 \ x\ if \0 \ x * 2 ^ n + y mod 2 ^ n\ proof - have \y mod 2 ^ n < 2 ^ n\ using pos_mod_bound [of \2 ^ n\ y] by simp then have \\ y mod 2 ^ n \ 2 ^ n\ by (simp add: less_le) with that have \x \ - 1\ by auto have *: \- 1 \ (- (y mod 2 ^ n)) div 2 ^ n\ by (simp add: zdiv_zminus1_eq_if) from that have \- (y mod 2 ^ n) \ x * 2 ^ n\ by simp then have \(- (y mod 2 ^ n)) div 2 ^ n \ (x * 2 ^ n) div 2 ^ n\ using zdiv_mono1 zero_less_numeral zero_less_power by blast with * have \- 1 \ x * 2 ^ n div 2 ^ n\ by simp with \x \ - 1\ show ?thesis by simp qed then show ?thesis by (simp add: bin_sign_def not_le not_less bin_cat_eq_push_bit_add_take_bit push_bit_eq_mult take_bit_eq_mod) qed lemma bin_cat_assoc: "bin_cat (bin_cat x m y) n z = bin_cat x (m + n) (bin_cat y n z)" by (fact concat_bit_assoc) lemma bin_cat_assoc_sym: "bin_cat x m (bin_cat y n z) = bin_cat (bin_cat x (m - n) y) (min m n) z" by (fact concat_bit_assoc_sym) definition bin_rcat :: "nat \ int list \ int" where "bin_rcat n = foldl (\u v. bin_cat u n v) 0" fun bin_rsplit_aux :: "nat \ nat \ int \ int list \ int list" where "bin_rsplit_aux n m c bs = (if m = 0 \ n = 0 then bs else let (a, b) = bin_split n c in bin_rsplit_aux n (m - n) a (b # bs))" definition bin_rsplit :: "nat \ nat \ int \ int list" where "bin_rsplit n w = bin_rsplit_aux n (fst w) (snd w) []" fun bin_rsplitl_aux :: "nat \ nat \ int \ int list \ int list" where "bin_rsplitl_aux n m c bs = (if m = 0 \ n = 0 then bs else let (a, b) = bin_split (min m n) c in bin_rsplitl_aux n (m - n) a (b # bs))" definition bin_rsplitl :: "nat \ nat \ int \ int list" where "bin_rsplitl n w = bin_rsplitl_aux n (fst w) (snd w) []" declare bin_rsplit_aux.simps [simp del] declare bin_rsplitl_aux.simps [simp del] lemma bin_nth_cat: "bin_nth (bin_cat x k y) n = (if n < k then bin_nth y n else bin_nth x (n - k))" by (simp add: bit_concat_bit_iff) lemma bin_nth_drop_bit_iff: \bin_nth (drop_bit n c) k \ bin_nth c (n + k)\ by (simp add: bit_drop_bit_eq) lemma bin_nth_take_bit_iff: \bin_nth (take_bit n c) k \ k < n \ bin_nth c k\ by (fact bit_take_bit_iff) lemma bin_nth_split: "bin_split n c = (a, b) \ (\k. bin_nth a k = bin_nth c (n + k)) \ (\k. bin_nth b k = (k < n \ bin_nth c k))" by (auto simp add: bin_nth_drop_bit_iff bin_nth_take_bit_iff) lemma bin_cat_zero [simp]: "bin_cat 0 n w = bintrunc n w" by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma bintr_cat1: "bintrunc (k + n) (bin_cat a n b) = bin_cat (bintrunc k a) n b" by (metis bin_cat_assoc bin_cat_zero) lemma bintr_cat: "bintrunc m (bin_cat a n b) = bin_cat (bintrunc (m - n) a) n (bintrunc (min m n) b)" by (rule bin_eqI) (auto simp: bin_nth_cat nth_bintr) lemma bintr_cat_same [simp]: "bintrunc n (bin_cat a n b) = bintrunc n b" by (auto simp add : bintr_cat) lemma cat_bintr [simp]: "bin_cat a n (bintrunc n b) = bin_cat a n b" by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma split_bintrunc: "bin_split n c = (a, b) \ b = bintrunc n c" by simp lemma bin_cat_split: "bin_split n w = (u, v) \ w = bin_cat u n v" by (auto simp add: bin_cat_eq_push_bit_add_take_bit bits_ident) lemma drop_bit_bin_cat_eq: \drop_bit n (bin_cat v n w) = v\ by (rule bit_eqI) (simp add: bit_drop_bit_eq bit_concat_bit_iff) lemma take_bit_bin_cat_eq: \take_bit n (bin_cat v n w) = take_bit n w\ by (rule bit_eqI) (simp add: bit_concat_bit_iff) lemma bin_split_cat: "bin_split n (bin_cat v n w) = (v, bintrunc n w)" by (simp add: drop_bit_bin_cat_eq take_bit_bin_cat_eq) lemma bin_split_zero [simp]: "bin_split n 0 = (0, 0)" by simp lemma bin_split_minus1 [simp]: "bin_split n (- 1) = (- 1, bintrunc n (- 1))" by simp lemma bin_split_trunc: "bin_split (min m n) c = (a, b) \ bin_split n (bintrunc m c) = (bintrunc (m - n) a, b)" apply (induct n arbitrary: m b c, clarsimp) apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) apply (case_tac m) apply (auto simp: Let_def drop_bit_Suc take_bit_Suc mod_2_eq_odd split: prod.split_asm) done lemma bin_split_trunc1: "bin_split n c = (a, b) \ bin_split n (bintrunc m c) = (bintrunc (m - n) a, bintrunc m b)" apply (induct n arbitrary: m b c, clarsimp) apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) apply (case_tac m) apply (auto simp: Let_def drop_bit_Suc take_bit_Suc mod_2_eq_odd split: prod.split_asm) done lemma bin_cat_num: "bin_cat a n b = a * 2 ^ n + bintrunc n b" by (simp add: bin_cat_eq_push_bit_add_take_bit push_bit_eq_mult) lemma bin_split_num: "bin_split n b = (b div 2 ^ n, b mod 2 ^ n)" by (simp add: drop_bit_eq_div take_bit_eq_mod) lemmas bin_rsplit_aux_simps = bin_rsplit_aux.simps bin_rsplitl_aux.simps lemmas rsplit_aux_simps = bin_rsplit_aux_simps lemmas th_if_simp1 = if_split [where P = "(=) l", THEN iffD1, THEN conjunct1, THEN mp] for l lemmas th_if_simp2 = if_split [where P = "(=) l", THEN iffD1, THEN conjunct2, THEN mp] for l lemmas rsplit_aux_simp1s = rsplit_aux_simps [THEN th_if_simp1] lemmas rsplit_aux_simp2ls = rsplit_aux_simps [THEN th_if_simp2] \ \these safe to \[simp add]\ as require calculating \m - n\\ lemmas bin_rsplit_aux_simp2s [simp] = rsplit_aux_simp2ls [unfolded Let_def] lemmas rbscl = bin_rsplit_aux_simp2s (2) lemmas rsplit_aux_0_simps [simp] = rsplit_aux_simp1s [OF disjI1] rsplit_aux_simp1s [OF disjI2] lemma bin_rsplit_aux_append: "bin_rsplit_aux n m c (bs @ cs) = bin_rsplit_aux n m c bs @ cs" apply (induct n m c bs rule: bin_rsplit_aux.induct) apply (subst bin_rsplit_aux.simps) apply (subst bin_rsplit_aux.simps) apply (clarsimp split: prod.split) done lemma bin_rsplitl_aux_append: "bin_rsplitl_aux n m c (bs @ cs) = bin_rsplitl_aux n m c bs @ cs" apply (induct n m c bs rule: bin_rsplitl_aux.induct) apply (subst bin_rsplitl_aux.simps) apply (subst bin_rsplitl_aux.simps) apply (clarsimp split: prod.split) done lemmas rsplit_aux_apps [where bs = "[]"] = bin_rsplit_aux_append bin_rsplitl_aux_append lemmas rsplit_def_auxs = bin_rsplit_def bin_rsplitl_def lemmas rsplit_aux_alts = rsplit_aux_apps [unfolded append_Nil rsplit_def_auxs [symmetric]] lemma bin_split_minus: "0 < n \ bin_split (Suc (n - 1)) w = bin_split n w" by auto lemma bin_split_pred_simp [simp]: "(0::nat) < numeral bin \ bin_split (numeral bin) w = (let (w1, w2) = bin_split (numeral bin - 1) (bin_rest w) in (w1, of_bool (odd w) + 2 * w2))" by (simp add: take_bit_rec drop_bit_rec mod_2_eq_odd) lemma bin_rsplit_aux_simp_alt: "bin_rsplit_aux n m c bs = (if m = 0 \ n = 0 then bs else let (a, b) = bin_split n c in bin_rsplit n (m - n, a) @ b # bs)" apply (simp add: bin_rsplit_aux.simps [of n m c bs]) apply (subst rsplit_aux_alts) apply (simp add: bin_rsplit_def) done lemmas bin_rsplit_simp_alt = trans [OF bin_rsplit_def bin_rsplit_aux_simp_alt] lemmas bthrs = bin_rsplit_simp_alt [THEN [2] trans] lemma bin_rsplit_size_sign' [rule_format]: "n > 0 \ rev sw = bin_rsplit n (nw, w) \ \v\set sw. bintrunc n v = v" apply (induct sw arbitrary: nw w) apply clarsimp apply clarsimp apply (drule bthrs) apply (simp (no_asm_use) add: Let_def split: prod.split_asm if_split_asm) apply clarify apply simp done lemmas bin_rsplit_size_sign = bin_rsplit_size_sign' [OF asm_rl rev_rev_ident [THEN trans] set_rev [THEN equalityD2 [THEN subsetD]]] lemma bin_nth_rsplit [rule_format] : "n > 0 \ m < n \ \w k nw. rev sw = bin_rsplit n (nw, w) \ k < size sw \ bin_nth (sw ! k) m = bin_nth w (k * n + m)" apply (induct sw) apply clarsimp apply clarsimp apply (drule bthrs) apply (simp (no_asm_use) add: Let_def split: prod.split_asm if_split_asm) apply (erule allE, erule impE, erule exI) apply (case_tac k) apply clarsimp prefer 2 apply clarsimp apply (erule allE) apply (erule (1) impE) apply (simp add: bit_drop_bit_eq ac_simps) apply (simp add: bit_take_bit_iff ac_simps) done lemma bin_rsplit_all: "0 < nw \ nw \ n \ bin_rsplit n (nw, w) = [bintrunc n w]" by (auto simp: bin_rsplit_def rsplit_aux_simp2ls split: prod.split dest!: split_bintrunc) lemma bin_rsplit_l [rule_format]: "\bin. bin_rsplitl n (m, bin) = bin_rsplit n (m, bintrunc m bin)" apply (rule_tac a = "m" in wf_less_than [THEN wf_induct]) apply (simp (no_asm) add: bin_rsplitl_def bin_rsplit_def) apply (rule allI) apply (subst bin_rsplitl_aux.simps) apply (subst bin_rsplit_aux.simps) apply (clarsimp simp: Let_def split: prod.split) apply (simp add: ac_simps) apply (subst rsplit_aux_alts(1)) apply (subst rsplit_aux_alts(2)) apply clarsimp unfolding bin_rsplit_def bin_rsplitl_def apply (simp add: drop_bit_take_bit) apply (case_tac \x < n\) apply (simp_all add: not_less min_def) done lemma bin_rsplit_rcat [rule_format]: "n > 0 \ bin_rsplit n (n * size ws, bin_rcat n ws) = map (bintrunc n) ws" apply (unfold bin_rsplit_def bin_rcat_def) apply (rule_tac xs = ws in rev_induct) apply clarsimp apply clarsimp apply (subst rsplit_aux_alts) apply (simp add: drop_bit_bin_cat_eq take_bit_bin_cat_eq) done lemma bin_rsplit_aux_len_le [rule_format] : "\ws m. n \ 0 \ ws = bin_rsplit_aux n nw w bs \ length ws \ m \ nw + length bs * n \ m * n" proof - have *: R if d: "i \ j \ m < j'" and R1: "i * k \ j * k \ R" and R2: "Suc m * k' \ j' * k' \ R" for i j j' k k' m :: nat and R using d apply safe apply (rule R1, erule mult_le_mono1) apply (rule R2, erule Suc_le_eq [THEN iffD2 [THEN mult_le_mono1]]) done have **: "0 < sc \ sc - n + (n + lb * n) \ m * n \ sc + lb * n \ m * n" for sc m n lb :: nat apply safe apply arith apply (case_tac "sc \ n") apply arith apply (insert linorder_le_less_linear [of m lb]) apply (erule_tac k=n and k'=n in *) apply arith apply simp done show ?thesis apply (induct n nw w bs rule: bin_rsplit_aux.induct) apply (subst bin_rsplit_aux.simps) apply (simp add: ** Let_def split: prod.split) done qed lemma bin_rsplit_len_le: "n \ 0 \ ws = bin_rsplit n (nw, w) \ length ws \ m \ nw \ m * n" by (auto simp: bin_rsplit_def bin_rsplit_aux_len_le) lemma bin_rsplit_aux_len: "n \ 0 \ length (bin_rsplit_aux n nw w cs) = (nw + n - 1) div n + length cs" apply (induct n nw w cs rule: bin_rsplit_aux.induct) apply (subst bin_rsplit_aux.simps) apply (clarsimp simp: Let_def split: prod.split) apply (erule thin_rl) apply (case_tac m) apply simp apply (case_tac "m \ n") apply (auto simp add: div_add_self2) done lemma bin_rsplit_len: "n \ 0 \ length (bin_rsplit n (nw, w)) = (nw + n - 1) div n" by (auto simp: bin_rsplit_def bin_rsplit_aux_len) lemma bin_rsplit_aux_len_indep: "n \ 0 \ length bs = length cs \ length (bin_rsplit_aux n nw v bs) = length (bin_rsplit_aux n nw w cs)" proof (induct n nw w cs arbitrary: v bs rule: bin_rsplit_aux.induct) case (1 n m w cs v bs) show ?case proof (cases "m = 0") case True with \length bs = length cs\ show ?thesis by simp next case False from "1.hyps" [of \bin_split n w\ \drop_bit n w\ \take_bit n w\] \m \ 0\ \n \ 0\ have hyp: "\v bs. length bs = Suc (length cs) \ length (bin_rsplit_aux n (m - n) v bs) = length (bin_rsplit_aux n (m - n) (drop_bit n w) (take_bit n w # cs))" using bin_rsplit_aux_len by fastforce from \length bs = length cs\ \n \ 0\ show ?thesis by (auto simp add: bin_rsplit_aux_simp_alt Let_def bin_rsplit_len split: prod.split) qed qed lemma bin_rsplit_len_indep: "n \ 0 \ length (bin_rsplit n (nw, v)) = length (bin_rsplit n (nw, w))" apply (unfold bin_rsplit_def) apply (simp (no_asm)) apply (erule bin_rsplit_aux_len_indep) apply (rule refl) done subsection \Logical operations\ primrec bin_sc :: "nat \ bool \ int \ int" where Z: "bin_sc 0 b w = of_bool b + 2 * bin_rest w" | Suc: "bin_sc (Suc n) b w = of_bool (odd w) + 2 * bin_sc n b (w div 2)" lemma bin_nth_sc [simp]: "bit (bin_sc n b w) n \ b" by (induction n arbitrary: w) (simp_all add: bit_Suc) lemma bin_sc_sc_same [simp]: "bin_sc n c (bin_sc n b w) = bin_sc n c w" by (induction n arbitrary: w) (simp_all add: bit_Suc) lemma bin_sc_sc_diff: "m \ n \ bin_sc m c (bin_sc n b w) = bin_sc n b (bin_sc m c w)" apply (induct n arbitrary: w m) apply (case_tac [!] m) apply auto done lemma bin_nth_sc_gen: "bin_nth (bin_sc n b w) m = (if m = n then b else bin_nth w m)" apply (induct n arbitrary: w m) apply (case_tac m; simp add: bit_Suc) apply (case_tac m; simp add: bit_Suc) done lemma bin_sc_eq: \bin_sc n False = unset_bit n\ \bin_sc n True = Bit_Operations.set_bit n\ by (simp_all add: fun_eq_iff bit_eq_iff) (simp_all add: bin_nth_sc_gen bit_set_bit_iff bit_unset_bit_iff) lemma bin_sc_nth [simp]: "bin_sc n (bin_nth w n) w = w" by (rule bit_eqI) (simp add: bin_nth_sc_gen) lemma bin_sign_sc [simp]: "bin_sign (bin_sc n b w) = bin_sign w" proof (induction n arbitrary: w) case 0 then show ?case by (auto simp add: bin_sign_def) (use bin_rest_ge_0 in fastforce) next case (Suc n) from Suc [of \w div 2\] show ?case by (auto simp add: bin_sign_def split: if_splits) qed lemma bin_sc_bintr [simp]: "bintrunc m (bin_sc n x (bintrunc m w)) = bintrunc m (bin_sc n x w)" apply (cases x) apply (simp_all add: bin_sc_eq bit_eq_iff) apply (auto simp add: bit_take_bit_iff bit_set_bit_iff bit_unset_bit_iff) done lemma bin_clr_le: "bin_sc n False w \ w" by (simp add: bin_sc_eq unset_bit_less_eq) lemma bin_set_ge: "bin_sc n True w \ w" by (simp add: bin_sc_eq set_bit_greater_eq) lemma bintr_bin_clr_le: "bintrunc n (bin_sc m False w) \ bintrunc n w" by (simp add: bin_sc_eq take_bit_unset_bit_eq unset_bit_less_eq) lemma bintr_bin_set_ge: "bintrunc n (bin_sc m True w) \ bintrunc n w" by (simp add: bin_sc_eq take_bit_set_bit_eq set_bit_greater_eq) lemma bin_sc_FP [simp]: "bin_sc n False 0 = 0" by (induct n) auto lemma bin_sc_TM [simp]: "bin_sc n True (- 1) = - 1" by (induct n) auto lemmas bin_sc_simps = bin_sc.Z bin_sc.Suc bin_sc_TM bin_sc_FP lemma bin_sc_minus: "0 < n \ bin_sc (Suc (n - 1)) b w = bin_sc n b w" by auto lemmas bin_sc_Suc_minus = trans [OF bin_sc_minus [symmetric] bin_sc.Suc] lemma bin_sc_numeral [simp]: "bin_sc (numeral k) b w = of_bool (odd w) + 2 * bin_sc (pred_numeral k) b (w div 2)" by (simp add: numeral_eq_Suc) instantiation int :: semiring_bit_syntax begin definition [iff]: "i !! n \ bin_nth i n" definition "shiftl x n = x * 2 ^ n" for x :: int definition "shiftr x n = x div 2 ^ n" for x :: int instance by standard (simp_all add: fun_eq_iff shiftl_int_def shiftr_int_def push_bit_eq_mult drop_bit_eq_div) end lemma shiftl_eq_push_bit: \k << n = push_bit n k\ for k :: int by (fact shiftl_eq_push_bit) lemma shiftr_eq_drop_bit: \k >> n = drop_bit n k\ for k :: int by (fact shiftr_eq_drop_bit) subsubsection \Basic simplification rules\ lemmas int_not_def = not_int_def lemma int_not_simps [simp]: "NOT (0::int) = -1" "NOT (1::int) = -2" "NOT (- 1::int) = 0" "NOT (numeral w::int) = - numeral (w + Num.One)" "NOT (- numeral (Num.Bit0 w)::int) = numeral (Num.BitM w)" "NOT (- numeral (Num.Bit1 w)::int) = numeral (Num.Bit0 w)" by (simp_all add: not_int_def) lemma int_not_not: "NOT (NOT x) = x" for x :: int by (fact bit.double_compl) lemma int_and_0 [simp]: "0 AND x = 0" for x :: int by (fact bit.conj_zero_left) lemma int_and_m1 [simp]: "-1 AND x = x" for x :: int by (fact bit.conj_one_left) lemma int_or_zero [simp]: "0 OR x = x" for x :: int by (fact bit.disj_zero_left) lemma int_or_minus1 [simp]: "-1 OR x = -1" for x :: int by (fact bit.disj_one_left) lemma int_xor_zero [simp]: "0 XOR x = x" for x :: int by (fact bit.xor_zero_left) subsubsection \Binary destructors\ lemma bin_rest_NOT [simp]: "bin_rest (NOT x) = NOT (bin_rest x)" by (fact not_int_div_2) lemma bin_last_NOT [simp]: "bin_last (NOT x) \ \ bin_last x" by simp lemma bin_rest_AND [simp]: "bin_rest (x AND y) = bin_rest x AND bin_rest y" by (subst and_int_rec) auto lemma bin_last_AND [simp]: "bin_last (x AND y) \ bin_last x \ bin_last y" by (subst and_int_rec) auto lemma bin_rest_OR [simp]: "bin_rest (x OR y) = bin_rest x OR bin_rest y" by (subst or_int_rec) auto lemma bin_last_OR [simp]: "bin_last (x OR y) \ bin_last x \ bin_last y" by (subst or_int_rec) auto lemma bin_rest_XOR [simp]: "bin_rest (x XOR y) = bin_rest x XOR bin_rest y" by (subst xor_int_rec) auto lemma bin_last_XOR [simp]: "bin_last (x XOR y) \ (bin_last x \ bin_last y) \ \ (bin_last x \ bin_last y)" by (subst xor_int_rec) auto lemma bin_nth_ops: "\x y. bin_nth (x AND y) n \ bin_nth x n \ bin_nth y n" "\x y. bin_nth (x OR y) n \ bin_nth x n \ bin_nth y n" "\x y. bin_nth (x XOR y) n \ bin_nth x n \ bin_nth y n" "\x. bin_nth (NOT x) n \ \ bin_nth x n" by (simp_all add: bit_and_iff bit_or_iff bit_xor_iff bit_not_iff) subsubsection \Derived properties\ lemma int_xor_minus1 [simp]: "-1 XOR x = NOT x" for x :: int by (fact bit.xor_one_left) lemma int_xor_extra_simps [simp]: "w XOR 0 = w" "w XOR -1 = NOT w" for w :: int by simp_all lemma int_or_extra_simps [simp]: "w OR 0 = w" "w OR -1 = -1" for w :: int by simp_all lemma int_and_extra_simps [simp]: "w AND 0 = 0" "w AND -1 = w" for w :: int by simp_all text \Commutativity of the above.\ lemma bin_ops_comm: fixes x y :: int shows int_and_comm: "x AND y = y AND x" and int_or_comm: "x OR y = y OR x" and int_xor_comm: "x XOR y = y XOR x" by (simp_all add: ac_simps) lemma bin_ops_same [simp]: "x AND x = x" "x OR x = x" "x XOR x = 0" for x :: int by simp_all lemmas bin_log_esimps = int_and_extra_simps int_or_extra_simps int_xor_extra_simps int_and_0 int_and_m1 int_or_zero int_or_minus1 int_xor_zero int_xor_minus1 subsubsection \Basic properties of logical (bit-wise) operations\ lemma bbw_ao_absorb: "x AND (y OR x) = x \ x OR (y AND x) = x" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_ao_absorbs_other: "x AND (x OR y) = x \ (y AND x) OR x = x" "(y OR x) AND x = x \ x OR (x AND y) = x" "(x OR y) AND x = x \ (x AND y) OR x = x" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemmas bbw_ao_absorbs [simp] = bbw_ao_absorb bbw_ao_absorbs_other lemma int_xor_not: "(NOT x) XOR y = NOT (x XOR y) \ x XOR (NOT y) = NOT (x XOR y)" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_and_assoc: "(x AND y) AND z = x AND (y AND z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_or_assoc: "(x OR y) OR z = x OR (y OR z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_xor_assoc: "(x XOR y) XOR z = x XOR (y XOR z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemmas bbw_assocs = int_and_assoc int_or_assoc int_xor_assoc (* BH: Why are these declared as simp rules??? *) lemma bbw_lcs [simp]: "y AND (x AND z) = x AND (y AND z)" "y OR (x OR z) = x OR (y OR z)" "y XOR (x XOR z) = x XOR (y XOR z)" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_not_dist: "NOT (x OR y) = (NOT x) AND (NOT y)" "NOT (x AND y) = (NOT x) OR (NOT y)" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_oa_dist: "(x AND y) OR z = (x OR z) AND (y OR z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_ao_dist: "(x OR y) AND z = (x AND z) OR (y AND z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) (* Why were these declared simp??? declare bin_ops_comm [simp] bbw_assocs [simp] *) subsubsection \Simplification with numerals\ text \Cases for \0\ and \-1\ are already covered by other simp rules.\ lemma bin_rest_neg_numeral_BitM [simp]: "bin_rest (- numeral (Num.BitM w)) = - numeral w" by simp lemma bin_last_neg_numeral_BitM [simp]: "bin_last (- numeral (Num.BitM w))" by simp (* FIXME: The rule sets below are very large (24 rules for each operator). Is there a simpler way to do this? *) lemma int_and_numerals [simp]: "numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)" "numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND numeral y)" "numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)" "numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND numeral y)" "numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)" "numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND - numeral (y + Num.One))" "numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)" "numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND - numeral (y + Num.One))" "- numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND numeral y)" "- numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND numeral y)" "- numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND numeral y)" "- numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND numeral y)" "- numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND - numeral y)" "- numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND - numeral (y + Num.One))" "- numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND - numeral y)" "- numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND - numeral (y + Num.One))" "(1::int) AND numeral (Num.Bit0 y) = 0" "(1::int) AND numeral (Num.Bit1 y) = 1" "(1::int) AND - numeral (Num.Bit0 y) = 0" "(1::int) AND - numeral (Num.Bit1 y) = 1" "numeral (Num.Bit0 x) AND (1::int) = 0" "numeral (Num.Bit1 x) AND (1::int) = 1" "- numeral (Num.Bit0 x) AND (1::int) = 0" "- numeral (Num.Bit1 x) AND (1::int) = 1" by (rule bin_rl_eqI; simp)+ lemma int_or_numerals [simp]: "numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR numeral y)" "numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)" "numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR numeral y)" "numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)" "numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR - numeral y)" "numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))" "numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR - numeral y)" "numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))" "- numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR numeral y)" "- numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR numeral y)" "- numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)" "- numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)" "- numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR - numeral y)" "- numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR - numeral (y + Num.One))" "- numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral y)" "- numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral (y + Num.One))" "(1::int) OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)" "(1::int) OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)" "(1::int) OR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)" "(1::int) OR - numeral (Num.Bit1 y) = - numeral (Num.Bit1 y)" "numeral (Num.Bit0 x) OR (1::int) = numeral (Num.Bit1 x)" "numeral (Num.Bit1 x) OR (1::int) = numeral (Num.Bit1 x)" "- numeral (Num.Bit0 x) OR (1::int) = - numeral (Num.BitM x)" "- numeral (Num.Bit1 x) OR (1::int) = - numeral (Num.Bit1 x)" by (rule bin_rl_eqI; simp)+ lemma int_xor_numerals [simp]: "numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR numeral y)" "numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR numeral y)" "numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR numeral y)" "numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR numeral y)" "numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR - numeral y)" "numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR - numeral (y + Num.One))" "numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR - numeral y)" "numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR - numeral (y + Num.One))" "- numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR numeral y)" "- numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR numeral y)" "- numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR numeral y)" "- numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR numeral y)" "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR - numeral y)" "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR - numeral (y + Num.One))" "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR - numeral y)" "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR - numeral (y + Num.One))" "(1::int) XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)" "(1::int) XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)" "(1::int) XOR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)" "(1::int) XOR - numeral (Num.Bit1 y) = - numeral (Num.Bit0 (y + Num.One))" "numeral (Num.Bit0 x) XOR (1::int) = numeral (Num.Bit1 x)" "numeral (Num.Bit1 x) XOR (1::int) = numeral (Num.Bit0 x)" "- numeral (Num.Bit0 x) XOR (1::int) = - numeral (Num.BitM x)" "- numeral (Num.Bit1 x) XOR (1::int) = - numeral (Num.Bit0 (x + Num.One))" by (rule bin_rl_eqI; simp)+ subsubsection \Interactions with arithmetic\ lemma plus_and_or: "(x AND y) + (x OR y) = x + y" for x y :: int proof (induction x arbitrary: y rule: int_bit_induct) case zero then show ?case by simp next case minus then show ?case by simp next case (even x) from even.IH [of \y div 2\] show ?case by (auto simp add: and_int_rec [of _ y] or_int_rec [of _ y] elim: oddE) next case (odd x) from odd.IH [of \y div 2\] show ?case by (auto simp add: and_int_rec [of _ y] or_int_rec [of _ y] elim: oddE) qed lemma le_int_or: "bin_sign y = 0 \ x \ x OR y" for x y :: int by (simp add: bin_sign_def or_greater_eq split: if_splits) lemmas int_and_le = xtrans(3) [OF bbw_ao_absorbs (2) [THEN conjunct2, symmetric] le_int_or] text \Interaction between bit-wise and arithmetic: good example of \bin_induction\.\ lemma bin_add_not: "x + NOT x = (-1::int)" by (simp add: not_int_def) lemma AND_mod: "x AND (2 ^ n - 1) = x mod 2 ^ n" for x :: int by (simp flip: take_bit_eq_mod add: take_bit_eq_mask mask_eq_exp_minus_1) subsubsection \Comparison\ lemma AND_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" shows "0 \ x AND y" using assms by simp lemma OR_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" "0 \ y" shows "0 \ x OR y" using assms by simp lemma XOR_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" "0 \ y" shows "0 \ x XOR y" using assms by simp lemma AND_upper1 [simp]: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" shows "x AND y \ x" using assms by (induction x arbitrary: y rule: int_bit_induct) (simp_all add: and_int_rec [of \_ * 2\] and_int_rec [of \1 + _ * 2\] add_increasing) lemmas AND_upper1' [simp] = order_trans [OF AND_upper1] \<^marker>\contributor \Stefan Berghofer\\ lemmas AND_upper1'' [simp] = order_le_less_trans [OF AND_upper1] \<^marker>\contributor \Stefan Berghofer\\ lemma AND_upper2 [simp]: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ y" shows "x AND y \ y" using assms AND_upper1 [of y x] by (simp add: ac_simps) lemmas AND_upper2' [simp] = order_trans [OF AND_upper2] \<^marker>\contributor \Stefan Berghofer\\ lemmas AND_upper2'' [simp] = order_le_less_trans [OF AND_upper2] \<^marker>\contributor \Stefan Berghofer\\ lemma OR_upper: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" "x < 2 ^ n" "y < 2 ^ n" shows "x OR y < 2 ^ n" using assms proof (induction x arbitrary: y n rule: int_bit_induct) case zero then show ?case by simp next case minus then show ?case by simp next case (even x) from even.IH [of \n - 1\ \y div 2\] even.prems even.hyps show ?case by (cases n) (auto simp add: or_int_rec [of \_ * 2\] elim: oddE) next case (odd x) from odd.IH [of \n - 1\ \y div 2\] odd.prems odd.hyps show ?case by (cases n) (auto simp add: or_int_rec [of \1 + _ * 2\], linarith) qed lemma XOR_upper: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" "x < 2 ^ n" "y < 2 ^ n" shows "x XOR y < 2 ^ n" using assms proof (induction x arbitrary: y n rule: int_bit_induct) case zero then show ?case by simp next case minus then show ?case by simp next case (even x) from even.IH [of \n - 1\ \y div 2\] even.prems even.hyps show ?case by (cases n) (auto simp add: xor_int_rec [of \_ * 2\] elim: oddE) next case (odd x) from odd.IH [of \n - 1\ \y div 2\] odd.prems odd.hyps show ?case by (cases n) (auto simp add: xor_int_rec [of \1 + _ * 2\]) qed subsubsection \Truncating results of bit-wise operations\ lemma bin_trunc_ao: "bintrunc n x AND bintrunc n y = bintrunc n (x AND y)" "bintrunc n x OR bintrunc n y = bintrunc n (x OR y)" by (auto simp add: bin_eq_iff bin_nth_ops nth_bintr) lemma bin_trunc_xor: "bintrunc n (bintrunc n x XOR bintrunc n y) = bintrunc n (x XOR y)" by (auto simp add: bin_eq_iff bin_nth_ops nth_bintr) lemma bin_trunc_not: "bintrunc n (NOT (bintrunc n x)) = bintrunc n (NOT x)" by (auto simp add: bin_eq_iff bin_nth_ops nth_bintr) text \Want theorems of the form of \bin_trunc_xor\.\ lemma bintr_bintr_i: "x = bintrunc n y \ bintrunc n x = bintrunc n y" by auto lemmas bin_trunc_and = bin_trunc_ao(1) [THEN bintr_bintr_i] lemmas bin_trunc_or = bin_trunc_ao(2) [THEN bintr_bintr_i] subsubsection \More lemmas\ 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 (fact bit.conj_disj_distrib) 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 simp lemma int_nand_same_middle: fixes x :: int shows "x AND y AND NOT x = 0" by (simp add: bit_eq_iff bit_and_iff bit_not_iff) lemma and_xor_dist: fixes x :: int shows "x AND (y XOR z) = (x AND y) XOR (x AND z)" by (fact bit.conj_xor_distrib) lemma int_and_lt0 [simp]: \x AND y < 0 \ x < 0 \ y < 0\ for x y :: int by (fact and_negative_int_iff) lemma int_and_ge0 [simp]: \x AND y \ 0 \ x \ 0 \ y \ 0\ for x y :: int by (fact and_nonnegative_int_iff) lemma int_and_1: fixes x :: int shows "x AND 1 = x mod 2" by (fact and_one_eq) lemma int_1_and: fixes x :: int shows "1 AND x = x mod 2" by (fact one_and_eq) lemma int_or_lt0 [simp]: \x OR y < 0 \ x < 0 \ y < 0\ for x y :: int by (fact or_negative_int_iff) lemma int_or_ge0 [simp]: \x OR y \ 0 \ x \ 0 \ y \ 0\ for x y :: int by (fact or_nonnegative_int_iff) lemma int_xor_lt0 [simp]: \x XOR y < 0 \ (x < 0) \ (y < 0)\ for x y :: int by (fact xor_negative_int_iff) lemma int_xor_ge0 [simp]: \x XOR y \ 0 \ (x \ 0 \ y \ 0)\ for x y :: int by (fact xor_nonnegative_int_iff) lemma even_conv_AND: \even i \ i AND 1 = 0\ for i :: int by (simp add: and_one_eq mod2_eq_if) lemma bin_last_conv_AND: "bin_last i \ i AND 1 \ 0" by (simp add: and_one_eq mod2_eq_if) lemma bitval_bin_last: "of_bool (bin_last i) = i AND 1" by (simp add: and_one_eq mod2_eq_if) lemma bin_sign_and: "bin_sign (i AND j) = - (bin_sign i * bin_sign j)" by(simp add: bin_sign_def) lemma int_not_neg_numeral: "NOT (- numeral n) = (Num.sub n num.One :: int)" by(simp add: int_not_def) lemma int_neg_numeral_pOne_conv_not: "- numeral (n + num.One) = (NOT (numeral n) :: int)" by(simp add: int_not_def) subsection \Setting and clearing bits\ lemma int_shiftl_BIT: fixes x :: int shows int_shiftl0 [simp]: "x << 0 = x" and int_shiftl_Suc [simp]: "x << Suc n = 2 * (x << n)" by (auto simp add: shiftl_int_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)" by (simp add: bit_push_bit_iff_int shiftl_eq_push_bit) lemma bin_last_shiftr: "odd (x >> n) \ x !! n" for x :: int by (simp add: shiftr_eq_drop_bit bit_iff_odd_drop_bit) lemma bin_rest_shiftr [simp]: "bin_rest (x >> n) = x >> Suc n" by (simp add: bit_eq_iff shiftr_eq_drop_bit drop_bit_Suc bit_drop_bit_eq drop_bit_half) lemma bin_nth_shiftr [simp]: "bin_nth (x >> n) m = bin_nth x (n + m)" by (simp add: shiftr_eq_drop_bit bit_drop_bit_eq) lemma bin_nth_conv_AND: fixes x :: int shows "bin_nth x n \ x AND (1 << n) \ 0" by (simp add: bit_eq_iff) (auto simp add: shiftl_eq_push_bit bit_and_iff bit_push_bit_iff bit_exp_iff) 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 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" using int_shiftl_numeral [of Num.One w] by simp 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)" by 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" by (simp add: shiftr_eq_drop_bit) lemma int_shiftr_lt_0 [simp]: fixes i :: int shows "i >> n < 0 \ i < 0" by (metis int_shiftr_ge_0 not_less) 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 add: shiftr_eq_drop_bit numeral_eq_Suc add_One drop_bit_Suc) 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 add: shiftr_eq_drop_bit drop_bit_Suc add_One) 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" proof - from sign y x have \x \ 0\ and \y = 2 ^ n\ and \x < 2 ^ n\ by (simp_all add: bin_sign_def shiftl_eq_push_bit push_bit_eq_mult split: if_splits) from \0 \ x\ \x < 2 ^ n\ \m < n\ have \bit x m \ bit (x - 2 ^ n) m\ proof (induction m arbitrary: x n) case 0 then show ?case by simp next case (Suc m) moreover define q where \q = n - 1\ ultimately have n: \n = Suc q\ by simp have \(x - 2 ^ Suc q) div 2 = x div 2 - 2 ^ q\ by simp moreover from Suc.IH [of \x div 2\ q] Suc.prems have \bit (x div 2) m \ bit (x div 2 - 2 ^ q) m\ by (simp add: n) ultimately show ?case by (simp add: bit_Suc n) qed with \y = 2 ^ n\ show ?thesis by simp qed lemma bin_clr_conv_NAND: "bin_sc n False i = i AND NOT (1 << n)" by (induct n arbitrary: i) (rule bin_rl_eqI; simp)+ lemma bin_set_conv_OR: "bin_sc n True i = i OR (1 << n)" by (induct n arbitrary: i) (rule bin_rl_eqI; simp)+ - -subsection \Semantic interpretation of \<^typ>\bool list\ as \<^typ>\int\\ - -lemma bin_bl_bin': "bl_to_bin (bin_to_bl_aux n w bs) = bl_to_bin_aux bs (bintrunc n w)" - by (induct n arbitrary: w bs) (auto simp: bl_to_bin_def take_bit_Suc ac_simps mod_2_eq_odd) - -lemma bin_bl_bin [simp]: "bl_to_bin (bin_to_bl n w) = bintrunc n w" - by (auto simp: bin_to_bl_def bin_bl_bin') - -lemma bl_to_bin_rep_F: "bl_to_bin (replicate n False @ bl) = bl_to_bin bl" - by (simp add: bin_to_bl_zero_aux [symmetric] bin_bl_bin') (simp add: bl_to_bin_def) - -lemma bin_to_bl_trunc [simp]: "n \ m \ bin_to_bl n (bintrunc m w) = bin_to_bl n w" - by (auto intro: bl_to_bin_inj) - -lemma bin_to_bl_aux_bintr: - "bin_to_bl_aux n (bintrunc m bin) bl = - replicate (n - m) False @ bin_to_bl_aux (min n m) bin bl" - apply (induct n arbitrary: m bin bl) - apply clarsimp - apply clarsimp - apply (case_tac "m") - apply (clarsimp simp: bin_to_bl_zero_aux) - apply (erule thin_rl) - apply (induct_tac n) - apply (auto simp add: take_bit_Suc) - done - -lemma bin_to_bl_bintr: - "bin_to_bl n (bintrunc m bin) = replicate (n - m) False @ bin_to_bl (min n m) bin" - unfolding bin_to_bl_def by (rule bin_to_bl_aux_bintr) - -lemma bl_to_bin_rep_False: "bl_to_bin (replicate n False) = 0" - by (induct n) auto - -lemma len_bin_to_bl_aux: "length (bin_to_bl_aux n w bs) = n + length bs" - by (fact size_bin_to_bl_aux) - -lemma len_bin_to_bl: "length (bin_to_bl n w) = n" - by (fact size_bin_to_bl) (* FIXME: duplicate *) - -lemma sign_bl_bin': "bin_sign (bl_to_bin_aux bs w) = bin_sign w" - by (induction bs arbitrary: w) (simp_all add: bin_sign_def) - -lemma sign_bl_bin: "bin_sign (bl_to_bin bs) = 0" - by (simp add: bl_to_bin_def sign_bl_bin') - -lemma bl_sbin_sign_aux: "hd (bin_to_bl_aux (Suc n) w bs) = (bin_sign (sbintrunc n w) = -1)" - by (induction n arbitrary: w bs) (auto simp add: bin_sign_def even_iff_mod_2_eq_zero bit_Suc) - -lemma bl_sbin_sign: "hd (bin_to_bl (Suc n) w) = (bin_sign (sbintrunc n w) = -1)" - unfolding bin_to_bl_def by (rule bl_sbin_sign_aux) - -lemma bin_nth_of_bl_aux: - "bin_nth (bl_to_bin_aux bl w) n = - (n < size bl \ rev bl ! n \ n \ length bl \ bin_nth w (n - size bl))" - apply (induction bl arbitrary: w) - apply simp_all - apply safe - apply (simp_all add: not_le nth_append bit_double_iff even_bit_succ_iff split: if_splits) - done - -lemma bin_nth_of_bl: "bin_nth (bl_to_bin bl) n = (n < length bl \ rev bl ! n)" - by (simp add: bl_to_bin_def bin_nth_of_bl_aux) - -lemma bin_nth_bl: "n < m \ bin_nth w n = nth (rev (bin_to_bl m w)) n" - apply (induct n arbitrary: m w) - apply clarsimp - apply (case_tac m, clarsimp) - apply (clarsimp simp: bin_to_bl_def) - apply (simp add: bin_to_bl_aux_alt) - apply (case_tac m, clarsimp) - apply (clarsimp simp: bin_to_bl_def) - apply (simp add: bin_to_bl_aux_alt bit_Suc) - done - -lemma nth_bin_to_bl_aux: - "n < m + length bl \ (bin_to_bl_aux m w bl) ! n = - (if n < m then bit w (m - 1 - n) else bl ! (n - m))" - apply (induction bl arbitrary: w) - apply simp_all - apply (simp add: bin_nth_bl [of \m - Suc n\ m] rev_nth flip: bin_to_bl_def) - apply (metis One_nat_def Suc_pred add_diff_cancel_left' - add_diff_cancel_right' bin_to_bl_aux_alt bin_to_bl_def - diff_Suc_Suc diff_is_0_eq diff_zero less_Suc_eq_0_disj - less_antisym less_imp_Suc_add list.size(3) nat_less_le nth_append size_bin_to_bl_aux) - done - -lemma nth_bin_to_bl: "n < m \ (bin_to_bl m w) ! n = bin_nth w (m - Suc n)" - by (simp add: bin_to_bl_def nth_bin_to_bl_aux) - -lemma bl_to_bin_lt2p_aux: "bl_to_bin_aux bs w < (w + 1) * (2 ^ length bs)" -proof (induction bs arbitrary: w) - case Nil - then show ?case - by simp -next - case (Cons b bs) - from Cons.IH [of \1 + 2 * w\] Cons.IH [of \2 * w\] - show ?case - apply (auto simp add: algebra_simps) - apply (subst mult_2 [of \2 ^ length bs\]) - apply (simp only: add.assoc) - apply (rule pos_add_strict) - apply simp_all - done -qed - -lemma bl_to_bin_lt2p_drop: "bl_to_bin bs < 2 ^ length (dropWhile Not bs)" -proof (induct bs) - case Nil - then show ?case by simp -next - case (Cons b bs) - with bl_to_bin_lt2p_aux[where w=1] show ?case - by (simp add: bl_to_bin_def) -qed - -lemma bl_to_bin_lt2p: "bl_to_bin bs < 2 ^ length bs" - by (metis bin_bl_bin bintr_lt2p bl_bin_bl) - -lemma bl_to_bin_ge2p_aux: "bl_to_bin_aux bs w \ w * (2 ^ length bs)" -proof (induction bs arbitrary: w) - case Nil - then show ?case - by simp -next - case (Cons b bs) - from Cons.IH [of \1 + 2 * w\] Cons.IH [of \2 * w\] - show ?case - apply (auto simp add: algebra_simps) - apply (rule add_le_imp_le_left [of \2 ^ length bs\]) - apply (rule add_increasing) - apply simp_all - done -qed - -lemma bl_to_bin_ge0: "bl_to_bin bs \ 0" - apply (unfold bl_to_bin_def) - apply (rule xtrans(4)) - apply (rule bl_to_bin_ge2p_aux) - apply simp - done - -lemma butlast_rest_bin: "butlast (bin_to_bl n w) = bin_to_bl (n - 1) (bin_rest w)" - apply (unfold bin_to_bl_def) - apply (cases n, clarsimp) - apply clarsimp - apply (auto simp add: bin_to_bl_aux_alt) - done - -lemma butlast_bin_rest: "butlast bl = bin_to_bl (length bl - Suc 0) (bin_rest (bl_to_bin bl))" - using butlast_rest_bin [where w="bl_to_bin bl" and n="length bl"] by simp - -lemma butlast_rest_bl2bin_aux: - "bl \ [] \ bl_to_bin_aux (butlast bl) w = bin_rest (bl_to_bin_aux bl w)" - by (induct bl arbitrary: w) auto - -lemma butlast_rest_bl2bin: "bl_to_bin (butlast bl) = bin_rest (bl_to_bin bl)" - by (cases bl) (auto simp: bl_to_bin_def butlast_rest_bl2bin_aux) - -lemma trunc_bl2bin_aux: - "bintrunc m (bl_to_bin_aux bl w) = - bl_to_bin_aux (drop (length bl - m) bl) (bintrunc (m - length bl) w)" -proof (induct bl arbitrary: w) - case Nil - show ?case by simp -next - case (Cons b bl) - show ?case - proof (cases "m - length bl") - case 0 - then have "Suc (length bl) - m = Suc (length bl - m)" by simp - with Cons show ?thesis by simp - next - case (Suc n) - then have "m - Suc (length bl) = n" by simp - with Cons Suc show ?thesis by (simp add: take_bit_Suc ac_simps) - qed -qed - -lemma trunc_bl2bin: "bintrunc m (bl_to_bin bl) = bl_to_bin (drop (length bl - m) bl)" - by (simp add: bl_to_bin_def trunc_bl2bin_aux) - -lemma trunc_bl2bin_len [simp]: "bintrunc (length bl) (bl_to_bin bl) = bl_to_bin bl" - by (simp add: trunc_bl2bin) - -lemma bl2bin_drop: "bl_to_bin (drop k bl) = bintrunc (length bl - k) (bl_to_bin bl)" - apply (rule trans) - prefer 2 - apply (rule trunc_bl2bin [symmetric]) - apply (cases "k \ length bl") - apply auto - done - -lemma take_rest_power_bin: "m \ n \ take m (bin_to_bl n w) = bin_to_bl m ((bin_rest ^^ (n - m)) w)" - apply (rule nth_equalityI) - apply simp - apply (clarsimp simp add: nth_bin_to_bl nth_rest_power_bin) - done - -lemma last_bin_last': "size xs > 0 \ last xs \ bin_last (bl_to_bin_aux xs w)" - by (induct xs arbitrary: w) auto - -lemma last_bin_last: "size xs > 0 \ last xs \ bin_last (bl_to_bin xs)" - unfolding bl_to_bin_def by (erule last_bin_last') - -lemma bin_last_last: "bin_last w \ last (bin_to_bl (Suc n) w)" - by (simp add: bin_to_bl_def) (auto simp: bin_to_bl_aux_alt) - -lemma drop_bin2bl_aux: - "drop m (bin_to_bl_aux n bin bs) = - bin_to_bl_aux (n - m) bin (drop (m - n) bs)" - apply (induction n arbitrary: m bin bs) - apply auto - apply (case_tac "m \ n") - apply (auto simp add: not_le Suc_diff_le) - apply (case_tac "m - n") - apply auto - apply (use Suc_diff_Suc in fastforce) - done - -lemma drop_bin2bl: "drop m (bin_to_bl n bin) = bin_to_bl (n - m) bin" - by (simp add: bin_to_bl_def drop_bin2bl_aux) - -lemma take_bin2bl_lem1: "take m (bin_to_bl_aux m w bs) = bin_to_bl m w" - apply (induct m arbitrary: w bs) - apply clarsimp - apply clarsimp - apply (simp add: bin_to_bl_aux_alt) - apply (simp add: bin_to_bl_def) - apply (simp add: bin_to_bl_aux_alt) - done - -lemma take_bin2bl_lem: "take m (bin_to_bl_aux (m + n) w bs) = take m (bin_to_bl (m + n) w)" - by (induct n arbitrary: w bs) (simp_all (no_asm) add: bin_to_bl_def take_bin2bl_lem1, simp) - -lemma bin_split_take: "bin_split n c = (a, b) \ bin_to_bl m a = take m (bin_to_bl (m + n) c)" - apply (induct n arbitrary: b c) - apply clarsimp - apply (clarsimp simp: Let_def split: prod.split_asm) - apply (simp add: bin_to_bl_def) - apply (simp add: take_bin2bl_lem drop_bit_Suc) - done - -lemma bin_to_bl_drop_bit: - "k = m + n \ bin_to_bl m (drop_bit n c) = take m (bin_to_bl k c)" - using bin_split_take by simp - -lemma bin_split_take1: - "k = m + n \ bin_split n c = (a, b) \ bin_to_bl m a = take m (bin_to_bl k c)" - using bin_split_take by simp - -lemma bl_bin_bl_rep_drop: - "bin_to_bl n (bl_to_bin bl) = - replicate (n - length bl) False @ drop (length bl - n) bl" - by (simp add: bl_to_bin_inj bl_to_bin_rep_F trunc_bl2bin) - -lemma bl_to_bin_aux_cat: - "bl_to_bin_aux bs (bin_cat w nv v) = - bin_cat w (nv + length bs) (bl_to_bin_aux bs v)" - by (rule bit_eqI) - (auto simp add: bin_nth_of_bl_aux bin_nth_cat algebra_simps) - -lemma bin_to_bl_aux_cat: - "bin_to_bl_aux (nv + nw) (bin_cat v nw w) bs = - bin_to_bl_aux nv v (bin_to_bl_aux nw w bs)" - by (induction nw arbitrary: w bs) (simp_all add: concat_bit_Suc) - -lemma bl_to_bin_aux_alt: "bl_to_bin_aux bs w = bin_cat w (length bs) (bl_to_bin bs)" - using bl_to_bin_aux_cat [where nv = "0" and v = "0"] - by (simp add: bl_to_bin_def [symmetric]) - -lemma bin_to_bl_cat: - "bin_to_bl (nv + nw) (bin_cat v nw w) = - bin_to_bl_aux nv v (bin_to_bl nw w)" - by (simp add: bin_to_bl_def bin_to_bl_aux_cat) - -lemmas bl_to_bin_aux_app_cat = - trans [OF bl_to_bin_aux_append bl_to_bin_aux_alt] - -lemmas bin_to_bl_aux_cat_app = - trans [OF bin_to_bl_aux_cat bin_to_bl_aux_alt] - -lemma bl_to_bin_app_cat: - "bl_to_bin (bsa @ bs) = bin_cat (bl_to_bin bsa) (length bs) (bl_to_bin bs)" - by (simp only: bl_to_bin_aux_app_cat bl_to_bin_def) - -lemma bin_to_bl_cat_app: - "bin_to_bl (n + nw) (bin_cat w nw wa) = bin_to_bl n w @ bin_to_bl nw wa" - by (simp only: bin_to_bl_def bin_to_bl_aux_cat_app) - -text \\bl_to_bin_app_cat_alt\ and \bl_to_bin_app_cat\ are easily interderivable.\ -lemma bl_to_bin_app_cat_alt: "bin_cat (bl_to_bin cs) n w = bl_to_bin (cs @ bin_to_bl n w)" - by (simp add: bl_to_bin_app_cat) - -lemma mask_lem: "(bl_to_bin (True # replicate n False)) = bl_to_bin (replicate n True) + 1" - apply (unfold bl_to_bin_def) - apply (induct n) - apply simp - apply (simp only: Suc_eq_plus1 replicate_add append_Cons [symmetric] bl_to_bin_aux_append) - apply simp - done - -lemma bin_exhaust: - "(\x b. bin = of_bool b + 2 * x \ Q) \ Q" for bin :: int - apply (cases \even bin\) - apply (auto elim!: evenE oddE) - apply fastforce - apply fastforce - done - -primrec rbl_succ :: "bool list \ bool list" - where - Nil: "rbl_succ Nil = Nil" - | Cons: "rbl_succ (x # xs) = (if x then False # rbl_succ xs else True # xs)" - -primrec rbl_pred :: "bool list \ bool list" - where - Nil: "rbl_pred Nil = Nil" - | Cons: "rbl_pred (x # xs) = (if x then False # xs else True # rbl_pred xs)" - -primrec rbl_add :: "bool list \ bool list \ bool list" - where \ \result is length of first arg, second arg may be longer\ - Nil: "rbl_add Nil x = Nil" - | Cons: "rbl_add (y # ys) x = - (let ws = rbl_add ys (tl x) - in (y \ hd x) # (if hd x \ y then rbl_succ ws else ws))" - -primrec rbl_mult :: "bool list \ bool list \ bool list" - where \ \result is length of first arg, second arg may be longer\ - Nil: "rbl_mult Nil x = Nil" - | Cons: "rbl_mult (y # ys) x = - (let ws = False # rbl_mult ys x - in if y then rbl_add ws x else ws)" - -lemma size_rbl_pred: "length (rbl_pred bl) = length bl" - by (induct bl) auto - -lemma size_rbl_succ: "length (rbl_succ bl) = length bl" - by (induct bl) auto - -lemma size_rbl_add: "length (rbl_add bl cl) = length bl" - by (induct bl arbitrary: cl) (auto simp: Let_def size_rbl_succ) - -lemma size_rbl_mult: "length (rbl_mult bl cl) = length bl" - by (induct bl arbitrary: cl) (auto simp add: Let_def size_rbl_add) - -lemmas rbl_sizes [simp] = - size_rbl_pred size_rbl_succ size_rbl_add size_rbl_mult - -lemmas rbl_Nils = - rbl_pred.Nil rbl_succ.Nil rbl_add.Nil rbl_mult.Nil - -lemma rbl_add_app2: "length blb \ length bla \ rbl_add bla (blb @ blc) = rbl_add bla blb" - apply (induct bla arbitrary: blb) - apply simp - apply clarsimp - apply (case_tac blb, clarsimp) - apply (clarsimp simp: Let_def) - done - -lemma rbl_add_take2: - "length blb \ length bla \ rbl_add bla (take (length bla) blb) = rbl_add bla blb" - apply (induct bla arbitrary: blb) - apply simp - apply clarsimp - apply (case_tac blb, clarsimp) - apply (clarsimp simp: Let_def) - done - -lemma rbl_mult_app2: "length blb \ length bla \ rbl_mult bla (blb @ blc) = rbl_mult bla blb" - apply (induct bla arbitrary: blb) - apply simp - apply clarsimp - apply (case_tac blb, clarsimp) - apply (clarsimp simp: Let_def rbl_add_app2) - done - -lemma rbl_mult_take2: - "length blb \ length bla \ rbl_mult bla (take (length bla) blb) = rbl_mult bla blb" - apply (rule trans) - apply (rule rbl_mult_app2 [symmetric]) - apply simp - apply (rule_tac f = "rbl_mult bla" in arg_cong) - apply (rule append_take_drop_id) - done - -lemma rbl_add_split: - "P (rbl_add (y # ys) (x # xs)) = - (\ws. length ws = length ys \ ws = rbl_add ys xs \ - (y \ ((x \ P (False # rbl_succ ws)) \ (\ x \ P (True # ws)))) \ - (\ y \ P (x # ws)))" - by (cases y) (auto simp: Let_def) - -lemma rbl_mult_split: - "P (rbl_mult (y # ys) xs) = - (\ws. length ws = Suc (length ys) \ ws = False # rbl_mult ys xs \ - (y \ P (rbl_add ws xs)) \ (\ y \ P ws))" - by (auto simp: Let_def) - -lemma rbl_pred: "rbl_pred (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin - 1))" -proof (unfold bin_to_bl_def, induction n arbitrary: bin) - case 0 - then show ?case - by simp -next - case (Suc n) - obtain b k where \bin = of_bool b + 2 * k\ - using bin_exhaust by blast - moreover have \(2 * k - 1) div 2 = k - 1\ - using even_succ_div_2 [of \2 * (k - 1)\] - by simp - ultimately show ?case - using Suc [of \bin div 2\] - by simp (simp add: bin_to_bl_aux_alt) -qed - -lemma rbl_succ: "rbl_succ (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin + 1))" - apply (unfold bin_to_bl_def) - apply (induction n arbitrary: bin) - apply simp_all - apply (case_tac bin rule: bin_exhaust) - apply simp - apply (simp add: bin_to_bl_aux_alt ac_simps) - done - -lemma rbl_add: - "\bina binb. rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = - rev (bin_to_bl n (bina + binb))" - apply (unfold bin_to_bl_def) - apply (induct n) - apply simp - apply clarsimp - apply (case_tac bina rule: bin_exhaust) - apply (case_tac binb rule: bin_exhaust) - apply (case_tac b) - apply (case_tac [!] "ba") - apply (auto simp: rbl_succ bin_to_bl_aux_alt Let_def ac_simps) - done - -lemma rbl_add_long: - "m \ n \ rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = - rev (bin_to_bl n (bina + binb))" - apply (rule box_equals [OF _ rbl_add_take2 rbl_add]) - apply (rule_tac f = "rbl_add (rev (bin_to_bl n bina))" in arg_cong) - apply (rule rev_swap [THEN iffD1]) - apply (simp add: rev_take drop_bin2bl) - apply simp - done - -lemma rbl_mult_gt1: - "m \ length bl \ - rbl_mult bl (rev (bin_to_bl m binb)) = - rbl_mult bl (rev (bin_to_bl (length bl) binb))" - apply (rule trans) - apply (rule rbl_mult_take2 [symmetric]) - apply simp_all - apply (rule_tac f = "rbl_mult bl" in arg_cong) - apply (rule rev_swap [THEN iffD1]) - apply (simp add: rev_take drop_bin2bl) - done - -lemma rbl_mult_gt: - "m > n \ - rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = - rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb))" - by (auto intro: trans [OF rbl_mult_gt1]) - -lemmas rbl_mult_Suc = lessI [THEN rbl_mult_gt] - -lemma rbbl_Cons: "b # rev (bin_to_bl n x) = rev (bin_to_bl (Suc n) (of_bool b + 2 * x))" - by (simp add: bin_to_bl_def) (simp add: bin_to_bl_aux_alt) - -lemma rbl_mult: - "rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = - rev (bin_to_bl n (bina * binb))" - apply (induct n arbitrary: bina binb) - apply simp_all - apply (unfold bin_to_bl_def) - apply clarsimp - apply (case_tac bina rule: bin_exhaust) - apply (case_tac binb rule: bin_exhaust) - apply simp - apply (simp add: bin_to_bl_aux_alt) - apply (simp add: rbbl_Cons rbl_mult_Suc rbl_add algebra_simps) - done - -lemma sclem: "size (concat (map (bin_to_bl n) xs)) = length xs * n" - by (induct xs) auto - -lemma bin_cat_foldl_lem: - "foldl (\u. bin_cat u n) x xs = - bin_cat x (size xs * n) (foldl (\u. bin_cat u n) y xs)" - apply (induct xs arbitrary: x) - apply simp - apply (simp (no_asm)) - apply (frule asm_rl) - apply (drule meta_spec) - apply (erule trans) - apply (drule_tac x = "bin_cat y n a" in meta_spec) - apply (simp add: bin_cat_assoc_sym min.absorb2) - done - -lemma bin_rcat_bl: "bin_rcat n wl = bl_to_bin (concat (map (bin_to_bl n) wl))" - apply (unfold bin_rcat_def) - apply (rule sym) - apply (induct wl) - apply (auto simp add: bl_to_bin_append) - apply (simp add: bl_to_bin_aux_alt sclem) - apply (simp add: bin_cat_foldl_lem [symmetric]) - done - -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 bl_xor_aux_bin: - "map2 (\x y. x \ y) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = - bin_to_bl_aux n (v XOR w) (map2 (\x y. x \ y) bs cs)" - apply (induction n arbitrary: v w bs cs) - apply auto - apply (case_tac v rule: bin_exhaust) - apply (case_tac w rule: bin_exhaust) - apply clarsimp - done - -lemma bl_or_aux_bin: - "map2 (\) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = - bin_to_bl_aux n (v OR w) (map2 (\) bs cs)" - by (induct n arbitrary: v w bs cs) simp_all - -lemma bl_and_aux_bin: - "map2 (\) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = - bin_to_bl_aux n (v AND w) (map2 (\) bs cs)" - by (induction n arbitrary: v w bs cs) simp_all - -lemma bl_not_aux_bin: "map Not (bin_to_bl_aux n w cs) = bin_to_bl_aux n (NOT w) (map Not cs)" - by (induct n arbitrary: w cs) auto - -lemma bl_not_bin: "map Not (bin_to_bl n w) = bin_to_bl n (NOT w)" - by (simp add: bin_to_bl_def bl_not_aux_bin) - -lemma bl_and_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v AND w)" - by (simp add: bin_to_bl_def bl_and_aux_bin) - -lemma bl_or_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v OR w)" - by (simp add: bin_to_bl_def bl_or_aux_bin) - -lemma bl_xor_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v XOR w)" - using bl_xor_aux_bin by (simp add: bin_to_bl_def) - end