diff --git a/src/HOL/Library/Cardinality.thy b/src/HOL/Library/Cardinality.thy --- a/src/HOL/Library/Cardinality.thy +++ b/src/HOL/Library/Cardinality.thy @@ -1,533 +1,533 @@ (* Title: HOL/Library/Cardinality.thy Author: Brian Huffman, Andreas Lochbihler *) section \Cardinality of types\ theory Cardinality imports Phantom_Type begin subsection \Preliminary lemmas\ (* These should be moved elsewhere *) lemma (in type_definition) univ: "UNIV = Abs ` A" proof show "Abs ` A \ UNIV" by (rule subset_UNIV) show "UNIV \ Abs ` A" proof fix x :: 'b have "x = Abs (Rep x)" by (rule Rep_inverse [symmetric]) moreover have "Rep x \ A" by (rule Rep) ultimately show "x \ Abs ` A" by (rule image_eqI) qed qed lemma (in type_definition) card: "card (UNIV :: 'b set) = card A" by (simp add: univ card_image inj_on_def Abs_inject) subsection \Cardinalities of types\ syntax "_type_card" :: "type => nat" ("(1CARD/(1'(_')))") translations "CARD('t)" => "CONST card (CONST UNIV :: 't set)" print_translation \ let fun card_univ_tr' ctxt [Const (\<^const_syntax>\UNIV\, Type (_, [T]))] = Syntax.const \<^syntax_const>\_type_card\ $ Syntax_Phases.term_of_typ ctxt T in [(\<^const_syntax>\card\, card_univ_tr')] end \ lemma card_prod [simp]: "CARD('a \ 'b) = CARD('a) * CARD('b)" unfolding UNIV_Times_UNIV [symmetric] by (simp only: card_cartesian_product) lemma card_UNIV_sum: "CARD('a + 'b) = (if CARD('a) \ 0 \ CARD('b) \ 0 then CARD('a) + CARD('b) else 0)" unfolding UNIV_Plus_UNIV[symmetric] by(auto simp add: card_eq_0_iff card_Plus simp del: UNIV_Plus_UNIV) lemma card_sum [simp]: "CARD('a + 'b) = CARD('a::finite) + CARD('b::finite)" by(simp add: card_UNIV_sum) lemma card_UNIV_option: "CARD('a option) = (if CARD('a) = 0 then 0 else CARD('a) + 1)" proof - have "(None :: 'a option) \ range Some" by clarsimp thus ?thesis by (simp add: UNIV_option_conv card_eq_0_iff finite_range_Some card_image) qed lemma card_option [simp]: "CARD('a option) = Suc CARD('a::finite)" by(simp add: card_UNIV_option) lemma card_UNIV_set: "CARD('a set) = (if CARD('a) = 0 then 0 else 2 ^ CARD('a))" by(simp add: card_eq_0_iff card_Pow flip: Pow_UNIV) lemma card_set [simp]: "CARD('a set) = 2 ^ CARD('a::finite)" by(simp add: card_UNIV_set) lemma card_nat [simp]: "CARD(nat) = 0" by (simp add: card_eq_0_iff) lemma card_fun: "CARD('a \ 'b) = (if CARD('a) \ 0 \ CARD('b) \ 0 \ CARD('b) = 1 then CARD('b) ^ CARD('a) else 0)" proof - { assume "0 < CARD('a)" and "0 < CARD('b)" hence fina: "finite (UNIV :: 'a set)" and finb: "finite (UNIV :: 'b set)" by(simp_all only: card_ge_0_finite) from finite_distinct_list[OF finb] obtain bs where bs: "set bs = (UNIV :: 'b set)" and distb: "distinct bs" by blast from finite_distinct_list[OF fina] obtain as where as: "set as = (UNIV :: 'a set)" and dista: "distinct as" by blast have cb: "CARD('b) = length bs" unfolding bs[symmetric] distinct_card[OF distb] .. have ca: "CARD('a) = length as" unfolding as[symmetric] distinct_card[OF dista] .. let ?xs = "map (\ys. the \ map_of (zip as ys)) (List.n_lists (length as) bs)" have "UNIV = set ?xs" proof(rule UNIV_eq_I) fix f :: "'a \ 'b" from as have "f = the \ map_of (zip as (map f as))" by(auto simp add: map_of_zip_map) thus "f \ set ?xs" using bs by(auto simp add: set_n_lists) qed moreover have "distinct ?xs" unfolding distinct_map proof(intro conjI distinct_n_lists distb inj_onI) fix xs ys :: "'b list" assume xs: "xs \ set (List.n_lists (length as) bs)" and ys: "ys \ set (List.n_lists (length as) bs)" and eq: "the \ map_of (zip as xs) = the \ map_of (zip as ys)" from xs ys have [simp]: "length xs = length as" "length ys = length as" by(simp_all add: length_n_lists_elem) have "map_of (zip as xs) = map_of (zip as ys)" proof fix x from as bs have "\y. map_of (zip as xs) x = Some y" "\y. map_of (zip as ys) x = Some y" by(simp_all add: map_of_zip_is_Some[symmetric]) with eq show "map_of (zip as xs) x = map_of (zip as ys) x" by(auto dest: fun_cong[where x=x]) qed with dista show "xs = ys" by(simp add: map_of_zip_inject) qed hence "card (set ?xs) = length ?xs" by(simp only: distinct_card) moreover have "length ?xs = length bs ^ length as" by(simp add: length_n_lists) ultimately have "CARD('a \ 'b) = CARD('b) ^ CARD('a)" using cb ca by simp } moreover { assume cb: "CARD('b) = 1" then obtain b where b: "UNIV = {b :: 'b}" by(auto simp add: card_Suc_eq) have eq: "UNIV = {\x :: 'a. b ::'b}" proof(rule UNIV_eq_I) fix x :: "'a \ 'b" { fix y have "x y \ UNIV" .. hence "x y = b" unfolding b by simp } thus "x \ {\x. b}" by(auto) qed have "CARD('a \ 'b) = 1" unfolding eq by simp } ultimately show ?thesis by(auto simp del: One_nat_def)(auto simp add: card_eq_0_iff dest: finite_fun_UNIVD2 finite_fun_UNIVD1) qed corollary finite_UNIV_fun: "finite (UNIV :: ('a \ 'b) set) \ finite (UNIV :: 'a set) \ finite (UNIV :: 'b set) \ CARD('b) = 1" (is "?lhs \ ?rhs") proof - have "?lhs \ CARD('a \ 'b) > 0" by(simp add: card_gt_0_iff) also have "\ \ CARD('a) > 0 \ CARD('b) > 0 \ CARD('b) = 1" by(simp add: card_fun) also have "\ = ?rhs" by(simp add: card_gt_0_iff) finally show ?thesis . qed lemma card_literal: "CARD(String.literal) = 0" by(simp add: card_eq_0_iff infinite_literal) subsection \Classes with at least 1 and 2\ text \Class finite already captures "at least 1"\ lemma zero_less_card_finite [simp]: "0 < CARD('a::finite)" unfolding neq0_conv [symmetric] by simp lemma one_le_card_finite [simp]: "Suc 0 \ CARD('a::finite)" by (simp add: less_Suc_eq_le [symmetric]) class CARD_1 = assumes CARD_1: "CARD ('a) = 1" begin subclass finite proof from CARD_1 show "finite (UNIV :: 'a set)" using finite_UNIV_fun by fastforce qed end text \Class for cardinality "at least 2"\ class card2 = finite + assumes two_le_card: "2 \ CARD('a)" lemma one_less_card: "Suc 0 < CARD('a::card2)" using two_le_card [where 'a='a] by simp lemma one_less_int_card: "1 < int CARD('a::card2)" using one_less_card [where 'a='a] by simp subsection \A type class for deciding finiteness of types\ type_synonym 'a finite_UNIV = "('a, bool) phantom" class finite_UNIV = fixes finite_UNIV :: "('a, bool) phantom" assumes finite_UNIV: "finite_UNIV = Phantom('a) (finite (UNIV :: 'a set))" lemma finite_UNIV_code [code_unfold]: "finite (UNIV :: 'a :: finite_UNIV set) \ of_phantom (finite_UNIV :: 'a finite_UNIV)" by(simp add: finite_UNIV) subsection \A type class for computing the cardinality of types\ definition is_list_UNIV :: "'a list \ bool" where "is_list_UNIV xs = (let c = CARD('a) in if c = 0 then False else size (remdups xs) = c)" lemma is_list_UNIV_iff: "is_list_UNIV xs \ set xs = UNIV" by(auto simp add: is_list_UNIV_def Let_def card_eq_0_iff List.card_set[symmetric] dest: subst[where P="finite", OF _ finite_set] card_eq_UNIV_imp_eq_UNIV) type_synonym 'a card_UNIV = "('a, nat) phantom" class card_UNIV = finite_UNIV + fixes card_UNIV :: "'a card_UNIV" assumes card_UNIV: "card_UNIV = Phantom('a) CARD('a)" subsection \Instantiations for \card_UNIV\\ instantiation nat :: card_UNIV begin definition "finite_UNIV = Phantom(nat) False" definition "card_UNIV = Phantom(nat) 0" instance by intro_classes (simp_all add: finite_UNIV_nat_def card_UNIV_nat_def) end instantiation int :: card_UNIV begin definition "finite_UNIV = Phantom(int) False" definition "card_UNIV = Phantom(int) 0" -instance by intro_classes (simp_all add: card_UNIV_int_def finite_UNIV_int_def infinite_UNIV_int) +instance by intro_classes (simp_all add: card_UNIV_int_def finite_UNIV_int_def) end instantiation natural :: card_UNIV begin definition "finite_UNIV = Phantom(natural) False" definition "card_UNIV = Phantom(natural) 0" instance by standard (auto simp add: finite_UNIV_natural_def card_UNIV_natural_def card_eq_0_iff type_definition.univ [OF type_definition_natural] natural_eq_iff dest!: finite_imageD intro: inj_onI) end instantiation integer :: card_UNIV begin definition "finite_UNIV = Phantom(integer) False" definition "card_UNIV = Phantom(integer) 0" instance by standard (auto simp add: finite_UNIV_integer_def card_UNIV_integer_def card_eq_0_iff type_definition.univ [OF type_definition_integer] dest!: finite_imageD intro: inj_onI) end instantiation list :: (type) card_UNIV begin definition "finite_UNIV = Phantom('a list) False" definition "card_UNIV = Phantom('a list) 0" instance by intro_classes (simp_all add: card_UNIV_list_def finite_UNIV_list_def infinite_UNIV_listI) end instantiation unit :: card_UNIV begin definition "finite_UNIV = Phantom(unit) True" definition "card_UNIV = Phantom(unit) 1" instance by intro_classes (simp_all add: card_UNIV_unit_def finite_UNIV_unit_def) end instantiation bool :: card_UNIV begin definition "finite_UNIV = Phantom(bool) True" definition "card_UNIV = Phantom(bool) 2" instance by(intro_classes)(simp_all add: card_UNIV_bool_def finite_UNIV_bool_def) end instantiation char :: card_UNIV begin definition "finite_UNIV = Phantom(char) True" definition "card_UNIV = Phantom(char) 256" instance by intro_classes (simp_all add: card_UNIV_char_def card_UNIV_char finite_UNIV_char_def) end instantiation prod :: (finite_UNIV, finite_UNIV) finite_UNIV begin definition "finite_UNIV = Phantom('a \ 'b) (of_phantom (finite_UNIV :: 'a finite_UNIV) \ of_phantom (finite_UNIV :: 'b finite_UNIV))" instance by intro_classes (simp add: finite_UNIV_prod_def finite_UNIV finite_prod) end instantiation prod :: (card_UNIV, card_UNIV) card_UNIV begin definition "card_UNIV = Phantom('a \ 'b) (of_phantom (card_UNIV :: 'a card_UNIV) * of_phantom (card_UNIV :: 'b card_UNIV))" instance by intro_classes (simp add: card_UNIV_prod_def card_UNIV) end instantiation sum :: (finite_UNIV, finite_UNIV) finite_UNIV begin definition "finite_UNIV = Phantom('a + 'b) (of_phantom (finite_UNIV :: 'a finite_UNIV) \ of_phantom (finite_UNIV :: 'b finite_UNIV))" instance by intro_classes (simp add: finite_UNIV_sum_def finite_UNIV) end instantiation sum :: (card_UNIV, card_UNIV) card_UNIV begin definition "card_UNIV = Phantom('a + 'b) (let ca = of_phantom (card_UNIV :: 'a card_UNIV); cb = of_phantom (card_UNIV :: 'b card_UNIV) in if ca \ 0 \ cb \ 0 then ca + cb else 0)" instance by intro_classes (auto simp add: card_UNIV_sum_def card_UNIV card_UNIV_sum) end instantiation "fun" :: (finite_UNIV, card_UNIV) finite_UNIV begin definition "finite_UNIV = Phantom('a \ 'b) (let cb = of_phantom (card_UNIV :: 'b card_UNIV) in cb = 1 \ of_phantom (finite_UNIV :: 'a finite_UNIV) \ cb \ 0)" instance by intro_classes (auto simp add: finite_UNIV_fun_def Let_def card_UNIV finite_UNIV finite_UNIV_fun card_gt_0_iff) end instantiation "fun" :: (card_UNIV, card_UNIV) card_UNIV begin definition "card_UNIV = Phantom('a \ 'b) (let ca = of_phantom (card_UNIV :: 'a card_UNIV); cb = of_phantom (card_UNIV :: 'b card_UNIV) in if ca \ 0 \ cb \ 0 \ cb = 1 then cb ^ ca else 0)" instance by intro_classes (simp add: card_UNIV_fun_def card_UNIV Let_def card_fun) end instantiation option :: (finite_UNIV) finite_UNIV begin definition "finite_UNIV = Phantom('a option) (of_phantom (finite_UNIV :: 'a finite_UNIV))" instance by intro_classes (simp add: finite_UNIV_option_def finite_UNIV) end instantiation option :: (card_UNIV) card_UNIV begin definition "card_UNIV = Phantom('a option) (let c = of_phantom (card_UNIV :: 'a card_UNIV) in if c \ 0 then Suc c else 0)" instance by intro_classes (simp add: card_UNIV_option_def card_UNIV card_UNIV_option) end instantiation String.literal :: card_UNIV begin definition "finite_UNIV = Phantom(String.literal) False" definition "card_UNIV = Phantom(String.literal) 0" instance by intro_classes (simp_all add: card_UNIV_literal_def finite_UNIV_literal_def infinite_literal card_literal) end instantiation set :: (finite_UNIV) finite_UNIV begin definition "finite_UNIV = Phantom('a set) (of_phantom (finite_UNIV :: 'a finite_UNIV))" instance by intro_classes (simp add: finite_UNIV_set_def finite_UNIV Finite_Set.finite_set) end instantiation set :: (card_UNIV) card_UNIV begin definition "card_UNIV = Phantom('a set) (let c = of_phantom (card_UNIV :: 'a card_UNIV) in if c = 0 then 0 else 2 ^ c)" instance by intro_classes (simp add: card_UNIV_set_def card_UNIV_set card_UNIV) end lemma UNIV_finite_1: "UNIV = set [finite_1.a\<^sub>1]" by(auto intro: finite_1.exhaust) lemma UNIV_finite_2: "UNIV = set [finite_2.a\<^sub>1, finite_2.a\<^sub>2]" by(auto intro: finite_2.exhaust) lemma UNIV_finite_3: "UNIV = set [finite_3.a\<^sub>1, finite_3.a\<^sub>2, finite_3.a\<^sub>3]" by(auto intro: finite_3.exhaust) lemma UNIV_finite_4: "UNIV = set [finite_4.a\<^sub>1, finite_4.a\<^sub>2, finite_4.a\<^sub>3, finite_4.a\<^sub>4]" by(auto intro: finite_4.exhaust) lemma UNIV_finite_5: "UNIV = set [finite_5.a\<^sub>1, finite_5.a\<^sub>2, finite_5.a\<^sub>3, finite_5.a\<^sub>4, finite_5.a\<^sub>5]" by(auto intro: finite_5.exhaust) instantiation Enum.finite_1 :: card_UNIV begin definition "finite_UNIV = Phantom(Enum.finite_1) True" definition "card_UNIV = Phantom(Enum.finite_1) 1" instance by intro_classes (simp_all add: UNIV_finite_1 card_UNIV_finite_1_def finite_UNIV_finite_1_def) end instantiation Enum.finite_2 :: card_UNIV begin definition "finite_UNIV = Phantom(Enum.finite_2) True" definition "card_UNIV = Phantom(Enum.finite_2) 2" instance by intro_classes (simp_all add: UNIV_finite_2 card_UNIV_finite_2_def finite_UNIV_finite_2_def) end instantiation Enum.finite_3 :: card_UNIV begin definition "finite_UNIV = Phantom(Enum.finite_3) True" definition "card_UNIV = Phantom(Enum.finite_3) 3" instance by intro_classes (simp_all add: UNIV_finite_3 card_UNIV_finite_3_def finite_UNIV_finite_3_def) end instantiation Enum.finite_4 :: card_UNIV begin definition "finite_UNIV = Phantom(Enum.finite_4) True" definition "card_UNIV = Phantom(Enum.finite_4) 4" instance by intro_classes (simp_all add: UNIV_finite_4 card_UNIV_finite_4_def finite_UNIV_finite_4_def) end instantiation Enum.finite_5 :: card_UNIV begin definition "finite_UNIV = Phantom(Enum.finite_5) True" definition "card_UNIV = Phantom(Enum.finite_5) 5" instance by intro_classes (simp_all add: UNIV_finite_5 card_UNIV_finite_5_def finite_UNIV_finite_5_def) end subsection \Code setup for sets\ text \ Implement \<^term>\CARD('a)\ via \<^term>\card_UNIV\ and provide implementations for \<^term>\finite\, \<^term>\card\, \<^term>\(\)\, and \<^term>\(=)\if the calling context already provides \<^class>\finite_UNIV\ and \<^class>\card_UNIV\ instances. If we implemented the latter always via \<^term>\card_UNIV\, we would require instances of essentially all element types, i.e., a lot of instantiation proofs and -- at run time -- possibly slow dictionary constructions. \ context begin qualified definition card_UNIV' :: "'a card_UNIV" where [code del]: "card_UNIV' = Phantom('a) CARD('a)" lemma CARD_code [code_unfold]: "CARD('a) = of_phantom (card_UNIV' :: 'a card_UNIV)" by(simp add: card_UNIV'_def) lemma card_UNIV'_code [code]: "card_UNIV' = card_UNIV" by(simp add: card_UNIV card_UNIV'_def) end lemma card_Compl: "finite A \ card (- A) = card (UNIV :: 'a set) - card (A :: 'a set)" by (metis Compl_eq_Diff_UNIV card_Diff_subset top_greatest) context fixes xs :: "'a :: finite_UNIV list" begin qualified definition finite' :: "'a set \ bool" where [simp, code del, code_abbrev]: "finite' = finite" lemma finite'_code [code]: "finite' (set xs) \ True" "finite' (List.coset xs) \ of_phantom (finite_UNIV :: 'a finite_UNIV)" by(simp_all add: card_gt_0_iff finite_UNIV) end context fixes xs :: "'a :: card_UNIV list" begin qualified definition card' :: "'a set \ nat" where [simp, code del, code_abbrev]: "card' = card" lemma card'_code [code]: "card' (set xs) = length (remdups xs)" "card' (List.coset xs) = of_phantom (card_UNIV :: 'a card_UNIV) - length (remdups xs)" by(simp_all add: List.card_set card_Compl card_UNIV) qualified definition subset' :: "'a set \ 'a set \ bool" where [simp, code del, code_abbrev]: "subset' = (\)" lemma subset'_code [code]: "subset' A (List.coset ys) \ (\y \ set ys. y \ A)" "subset' (set ys) B \ (\y \ set ys. y \ B)" "subset' (List.coset xs) (set ys) \ (let n = CARD('a) in n > 0 \ card(set (xs @ ys)) = n)" by(auto simp add: Let_def card_gt_0_iff dest: card_eq_UNIV_imp_eq_UNIV intro: arg_cong[where f=card]) (metis finite_compl finite_set rev_finite_subset) qualified definition eq_set :: "'a set \ 'a set \ bool" where [simp, code del, code_abbrev]: "eq_set = (=)" lemma eq_set_code [code]: fixes ys defines "rhs \ let n = CARD('a) in if n = 0 then False else let xs' = remdups xs; ys' = remdups ys in length xs' + length ys' = n \ (\x \ set xs'. x \ set ys') \ (\y \ set ys'. y \ set xs')" shows "eq_set (List.coset xs) (set ys) \ rhs" and "eq_set (set ys) (List.coset xs) \ rhs" and "eq_set (set xs) (set ys) \ (\x \ set xs. x \ set ys) \ (\y \ set ys. y \ set xs)" and "eq_set (List.coset xs) (List.coset ys) \ (\x \ set xs. x \ set ys) \ (\y \ set ys. y \ set xs)" proof goal_cases { case 1 show ?case (is "?lhs \ ?rhs") proof show ?rhs if ?lhs using that by (auto simp add: rhs_def Let_def List.card_set[symmetric] card_Un_Int[where A="set xs" and B="- set xs"] card_UNIV Compl_partition card_gt_0_iff dest: sym)(metis finite_compl finite_set) show ?lhs if ?rhs proof - have "\ \y\set xs. y \ set ys; \x\set ys. x \ set xs \ \ set xs \ set ys = {}" by blast with that show ?thesis by (auto simp add: rhs_def Let_def List.card_set[symmetric] card_UNIV card_gt_0_iff card_Un_Int[where A="set xs" and B="set ys"] dest: card_eq_UNIV_imp_eq_UNIV split: if_split_asm) qed qed } moreover case 2 ultimately show ?case unfolding eq_set_def by blast next case 3 show ?case unfolding eq_set_def List.coset_def by blast next case 4 show ?case unfolding eq_set_def List.coset_def by blast qed end text \ Provide more informative exceptions than Match for non-rewritten cases. If generated code raises one these exceptions, then a code equation calls the mentioned operator for an element type that is not an instance of \<^class>\card_UNIV\ and is therefore not implemented via \<^term>\card_UNIV\. Constrain the element type with sort \<^class>\card_UNIV\ to change this. \ lemma card_coset_error [code]: "card (List.coset xs) = Code.abort (STR ''card (List.coset _) requires type class instance card_UNIV'') (\_. card (List.coset xs))" by(simp) lemma coset_subseteq_set_code [code]: "List.coset xs \ set ys \ (if xs = [] \ ys = [] then False else Code.abort (STR ''subset_eq (List.coset _) (List.set _) requires type class instance card_UNIV'') (\_. List.coset xs \ set ys))" by simp notepad begin \ \test code setup\ have "List.coset [True] = set [False] \ List.coset [] \ List.set [True, False] \ finite (List.coset [True])" by eval end end diff --git a/src/HOL/Library/Z2.thy b/src/HOL/Library/Z2.thy --- a/src/HOL/Library/Z2.thy +++ b/src/HOL/Library/Z2.thy @@ -1,180 +1,180 @@ (* Title: HOL/Library/Z2.thy Author: Brian Huffman *) section \The Field of Integers mod 2\ theory Z2 imports Main begin text \ Note that in most cases \<^typ>\bool\ is appropriate hen a binary type is needed; the type provided here, for historical reasons named \<^text>\bit\, is only needed if proper field operations are required. \ subsection \Bits as a datatype\ typedef bit = "UNIV :: bool set" morphisms set Bit .. instantiation bit :: "{zero, one}" begin definition zero_bit_def: "0 = Bit False" definition one_bit_def: "1 = Bit True" instance .. end old_rep_datatype "0::bit" "1::bit" proof - fix P :: "bit \ bool" fix x :: bit assume "P 0" and "P 1" then have "\b. P (Bit b)" unfolding zero_bit_def one_bit_def by (simp add: all_bool_eq) then show "P x" by (induct x) simp next show "(0::bit) \ (1::bit)" unfolding zero_bit_def one_bit_def by (simp add: Bit_inject) qed lemma Bit_set_eq [simp]: "Bit (set b) = b" by (fact set_inverse) lemma set_Bit_eq [simp]: "set (Bit P) = P" by (rule Bit_inverse) rule lemma bit_eq_iff: "x = y \ (set x \ set y)" by (auto simp add: set_inject) lemma Bit_inject [simp]: "Bit P = Bit Q \ (P \ Q)" by (auto simp add: Bit_inject) lemma set [iff]: "\ set 0" "set 1" by (simp_all add: zero_bit_def one_bit_def Bit_inverse) lemma [code]: "set 0 \ False" "set 1 \ True" by simp_all lemma set_iff: "set b \ b = 1" by (cases b) simp_all lemma bit_eq_iff_set: "b = 0 \ \ set b" "b = 1 \ set b" by (simp_all add: bit_eq_iff) lemma Bit [simp, code]: "Bit False = 0" "Bit True = 1" by (simp_all add: zero_bit_def one_bit_def) lemma bit_not_0_iff [iff]: "x \ 0 \ x = 1" for x :: bit by (simp add: bit_eq_iff) lemma bit_not_1_iff [iff]: "x \ 1 \ x = 0" for x :: bit by (simp add: bit_eq_iff) lemma [code]: "HOL.equal 0 b \ \ set b" "HOL.equal 1 b \ set b" by (simp_all add: equal set_iff) subsection \Type \<^typ>\bit\ forms a field\ instantiation bit :: field begin definition plus_bit_def: "x + y = case_bit y (case_bit 1 0 y) x" definition times_bit_def: "x * y = case_bit 0 y x" definition uminus_bit_def [simp]: "- x = x" for x :: bit definition minus_bit_def [simp]: "x - y = x + y" for x y :: bit definition inverse_bit_def [simp]: "inverse x = x" for x :: bit definition divide_bit_def [simp]: "x div y = x * y" for x y :: bit lemmas field_bit_defs = plus_bit_def times_bit_def minus_bit_def uminus_bit_def divide_bit_def inverse_bit_def instance - by standard (auto simp: field_bit_defs split: bit.split) + by standard (auto simp: plus_bit_def times_bit_def split: bit.split) end lemma bit_add_self: "x + x = 0" for x :: bit unfolding plus_bit_def by (simp split: bit.split) lemma bit_mult_eq_1_iff [simp]: "x * y = 1 \ x = 1 \ y = 1" for x y :: bit unfolding times_bit_def by (simp split: bit.split) text \Not sure whether the next two should be simp rules.\ lemma bit_add_eq_0_iff: "x + y = 0 \ x = y" for x y :: bit unfolding plus_bit_def by (simp split: bit.split) lemma bit_add_eq_1_iff: "x + y = 1 \ x \ y" for x y :: bit unfolding plus_bit_def by (simp split: bit.split) subsection \Numerals at type \<^typ>\bit\\ text \All numerals reduce to either 0 or 1.\ lemma bit_minus1 [simp]: "- 1 = (1 :: bit)" by (simp only: uminus_bit_def) lemma bit_neg_numeral [simp]: "(- numeral w :: bit) = numeral w" by (simp only: uminus_bit_def) lemma bit_numeral_even [simp]: "numeral (Num.Bit0 w) = (0 :: bit)" by (simp only: numeral_Bit0 bit_add_self) lemma bit_numeral_odd [simp]: "numeral (Num.Bit1 w) = (1 :: bit)" by (simp only: numeral_Bit1 bit_add_self add_0_left) subsection \Conversion from \<^typ>\bit\\ context zero_neq_one begin definition of_bit :: "bit \ 'a" where "of_bit b = case_bit 0 1 b" lemma of_bit_eq [simp, code]: "of_bit 0 = 0" "of_bit 1 = 1" by (simp_all add: of_bit_def) lemma of_bit_eq_iff: "of_bit x = of_bit y \ x = y" by (cases x) (cases y; simp)+ end lemma (in semiring_1) of_nat_of_bit_eq: "of_nat (of_bit b) = of_bit b" by (cases b) simp_all lemma (in ring_1) of_int_of_bit_eq: "of_int (of_bit b) = of_bit b" by (cases b) simp_all hide_const (open) set 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,2771 +1,2770 @@ (* Title: HOL/Word/Bits_Int.thy Author: Jeremy Dawson and Gerwin Klein, NICTA Definitions and basic theorems for bit-wise logical operations for integers expressed using Pls, Min, BIT, and converting them to and from lists of bools. *) section \Bitwise Operations on integers\ theory Bits_Int imports Bits Misc_Auxiliary begin subsection \Implicit bit representation of \<^typ>\int\\ definition Bit :: "int \ bool \ int" (infixl "BIT" 90) where "k BIT b = (if b then 1 else 0) + k + k" lemma Bit_B0: "k BIT False = k + k" by (simp add: Bit_def) lemma Bit_B1: "k BIT True = k + k + 1" by (simp add: Bit_def) lemma Bit_B0_2t: "k BIT False = 2 * k" by (rule trans, rule Bit_B0) simp lemma Bit_B1_2t: "k BIT True = 2 * k + 1" by (rule trans, rule Bit_B1) simp lemma uminus_Bit_eq: "- k BIT b = (- k - of_bool b) BIT b" by (cases b) (simp_all add: Bit_def) lemma power_BIT: "2 ^ Suc n - 1 = (2 ^ n - 1) BIT True" by (simp add: Bit_B1) 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) definition bin_rest :: "int \ int" where "bin_rest w = w div 2" lemma bin_rl_simp [simp]: "bin_rest w BIT bin_last w = w" unfolding bin_rest_def bin_last_def Bit_def by (cases "w mod 2 = 0") (use div_mult_mod_eq [of w 2] in simp_all) lemma bin_rest_BIT [simp]: "bin_rest (x BIT b) = x" unfolding bin_rest_def Bit_def by (cases b) simp_all lemma even_BIT [simp]: "even (x BIT b) \ \ b" by (simp add: Bit_def) lemma bin_last_BIT [simp]: "bin_last (x BIT b) = b" by simp lemma BIT_eq_iff [iff]: "u BIT b = v BIT c \ u = v \ b = c" by (auto simp: Bit_def) arith+ lemma BIT_bin_simps [simp]: "numeral k BIT False = numeral (Num.Bit0 k)" "numeral k BIT True = numeral (Num.Bit1 k)" "(- numeral k) BIT False = - numeral (Num.Bit0 k)" "(- numeral k) BIT True = - numeral (Num.BitM k)" - unfolding numeral.simps numeral_BitM - by (simp_all add: Bit_def del: arith_simps add_numeral_special diff_numeral_special) + by (simp_all only: Bit_B0 Bit_B1 numeral.simps numeral_BitM) lemma BIT_special_simps [simp]: shows "0 BIT False = 0" and "0 BIT True = 1" and "1 BIT False = 2" and "1 BIT True = 3" and "(- 1) BIT False = - 2" and "(- 1) BIT True = - 1" by (simp_all add: Bit_def) lemma Bit_eq_0_iff: "w BIT b = 0 \ w = 0 \ \ b" by (auto simp: Bit_def) arith lemma Bit_eq_m1_iff: "w BIT b = -1 \ w = -1 \ b" by (auto simp: Bit_def) arith lemma BitM_inc: "Num.BitM (Num.inc w) = Num.Bit1 w" by (induct w) simp_all lemma expand_BIT: "numeral (Num.Bit0 w) = numeral w BIT False" "numeral (Num.Bit1 w) = numeral w BIT True" "- numeral (Num.Bit0 w) = (- numeral w) BIT False" "- numeral (Num.Bit1 w) = (- numeral (w + Num.One)) BIT True" by (simp_all add: add_One BitM_inc) 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 add: bin_last_def zmod_zminus1_eq_if) 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 add: bin_rest_def zdiv_zminus1_eq_if) lemma less_Bits: "v BIT b < w BIT c \ v < w \ v \ w \ \ b \ c" by (auto simp: Bit_def) lemma le_Bits: "v BIT b \ w BIT c \ v < w \ v \ w \ (\ b \ c)" by (auto simp: Bit_def) lemma pred_BIT_simps [simp]: "x BIT False - 1 = (x - 1) BIT True" "x BIT True - 1 = x BIT False" by (simp_all add: Bit_B0_2t Bit_B1_2t) lemma succ_BIT_simps [simp]: "x BIT False + 1 = x BIT True" "x BIT True + 1 = (x + 1) BIT False" by (simp_all add: Bit_B0_2t Bit_B1_2t) lemma add_BIT_simps [simp]: "x BIT False + y BIT False = (x + y) BIT False" "x BIT False + y BIT True = (x + y) BIT True" "x BIT True + y BIT False = (x + y) BIT True" "x BIT True + y BIT True = (x + y + 1) BIT False" by (simp_all add: Bit_B0_2t Bit_B1_2t) lemma mult_BIT_simps [simp]: "x BIT False * y = (x * y) BIT False" "x * y BIT False = (x * y) BIT False" "x BIT True * y = (x * y) BIT False + y" by (simp_all add: Bit_B0_2t Bit_B1_2t algebra_simps) lemma B_mod_2': "X = 2 \ (w BIT True) mod X = 1 \ (w BIT False) mod X = 0" by (simp add: Bit_B0 Bit_B1) lemma bin_ex_rl: "\w b. w BIT b = bin" by (metis bin_rl_simp) lemma bin_exhaust: "(\x b. bin = x BIT b \ Q) \ Q" by (metis bin_ex_rl) lemma bin_abs_lem: "bin = (w BIT b) \ bin \ -1 \ bin \ 0 \ nat \w\ < nat \bin\" apply clarsimp apply (unfold Bit_def) apply (cases b) apply (clarsimp, arith) apply (clarsimp, arith) done lemma bin_induct: assumes PPls: "P 0" and PMin: "P (- 1)" and PBit: "\bin bit. P bin \ P (bin BIT bit)" shows "P bin" apply (rule_tac P=P and a=bin and f1="nat \ abs" in wf_measure [THEN wf_induct]) apply (simp add: measure_def inv_image_def) apply (case_tac x rule: bin_exhaust) apply (frule bin_abs_lem) apply (auto simp add : PPls PMin PBit) done lemma Bit_div2 [simp]: "(w BIT b) div 2 = w" unfolding bin_rest_def [symmetric] by (rule bin_rest_BIT) lemma bin_rl_eqI: "\bin_rest x = bin_rest y; bin_last x = bin_last y\ \ x = y" by (metis (mono_tags) BIT_eq_iff bin_ex_rl bin_last_BIT bin_rest_BIT) lemma twice_conv_BIT: "2 * x = x BIT False" by (rule bin_rl_eqI) (simp_all, simp_all add: bin_rest_def bin_last_def) lemma BIT_lt0 [simp]: "x BIT b < 0 \ x < 0" by(cases b)(auto simp add: Bit_def) lemma BIT_ge0 [simp]: "x BIT b \ 0 \ x \ 0" by(cases b)(auto simp add: Bit_def) lemma [simp]: shows bin_rest_lt0: "bin_rest i < 0 \ i < 0" and bin_rest_ge_0: "bin_rest i \ 0 \ i \ 0" by(auto simp add: bin_rest_def) lemma bin_rest_gt_0 [simp]: "bin_rest x > 0 \ x > 1" by(simp add: bin_rest_def add1_zle_eq pos_imp_zdiv_pos_iff) (metis add1_zle_eq one_add_one) 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 (w BIT b)" 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_Bit_minus_simp [simp]: "0 < n \ bin_to_bl_aux n (w BIT b) bl = bin_to_bl_aux (n - 1) w (b # 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 only: bin_to_bl_aux.simps bin_last_numeral_simps, simp) 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 only: bin_to_bl_aux.simps bin_last_numeral_simps, simp) 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) lemma bl_to_bin_BIT: "bl_to_bin bs BIT b = bl_to_bin (bs @ [b])" by (simp add: bl_to_bin_append) subsection \Bit projection\ primrec bin_nth :: "int \ nat \ bool" where Z: "bin_nth w 0 \ bin_last w" | Suc: "bin_nth w (Suc n) \ bin_nth (bin_rest w) n" lemma bin_nth_eq_mod: "bin_nth w n \ odd (w div 2 ^ n)" by (induction n arbitrary: w) (simp_all add: bin_last_def bin_rest_def odd_iff_mod_2_eq_one zdiv_zmult2_eq) lemma bin_nth_eq_iff: "bin_nth x = bin_nth y \ x = y" proof - have bin_nth_lem [rule_format]: "\y. bin_nth x = bin_nth y \ x = y" apply (induct x rule: bin_induct) apply safe apply (erule rev_mp) apply (induct_tac y rule: bin_induct) apply safe apply (drule_tac x=0 in fun_cong, force) apply (erule notE, rule ext, drule_tac x="Suc x" in fun_cong, force) apply (drule_tac x=0 in fun_cong, force) apply (erule rev_mp) apply (induct_tac y rule: bin_induct) apply safe apply (drule_tac x=0 in fun_cong, force) apply (erule notE, rule ext, drule_tac x="Suc x" in fun_cong, force) apply (metis Bit_eq_m1_iff Z bin_last_BIT) apply (case_tac y rule: bin_exhaust) apply clarify apply (erule allE) apply (erule impE) prefer 2 apply (erule conjI) apply (drule_tac x=0 in fun_cong, force) apply (rule ext) apply (drule_tac x="Suc x" for x in fun_cong, force) done show ?thesis by (auto elim: bin_nth_lem) qed 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)" using bin_nth_eq_iff by auto lemma bin_nth_zero [simp]: "\ bin_nth 0 n" by (induct n) auto lemma bin_nth_1 [simp]: "bin_nth 1 n \ n = 0" by (cases n) simp_all lemma bin_nth_minus1 [simp]: "bin_nth (- 1) n" by (induct n) auto lemma bin_nth_0_BIT: "bin_nth (w BIT b) 0 \ b" by auto lemma bin_nth_Suc_BIT: "bin_nth (w BIT b) (Suc n) = bin_nth w n" by auto lemma bin_nth_minus [simp]: "0 < n \ bin_nth (w BIT b) n = bin_nth w (n - 1)" by (cases n) auto 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) 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 = bin_nth.Z bin_nth.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\\ apply (induct n arbitrary: m) apply clarsimp apply safe apply (case_tac m) apply (auto simp: Bit_B0_2t [symmetric]) done 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: bin_nth.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" "bin_sign (w BIT b) = bin_sign w" by (simp_all add: bin_sign_def Bit_def) lemma bin_sign_rest [simp]: "bin_sign (bin_rest w) = bin_sign w" by (cases w rule: bin_exhaust) auto primrec bintrunc :: "nat \ int \ int" where Z : "bintrunc 0 bin = 0" | Suc : "bintrunc (Suc n) bin = bintrunc n (bin_rest bin) BIT (bin_last bin)" primrec sbintrunc :: "nat \ int \ int" where Z : "sbintrunc 0 bin = (if bin_last bin then -1 else 0)" | Suc : "sbintrunc (Suc n) bin = sbintrunc n (bin_rest bin) BIT (bin_last bin)" lemma bintrunc_mod2p: "bintrunc n w = w mod 2 ^ n" by (induct n arbitrary: w) (auto simp add: bin_last_def bin_rest_def Bit_def zmod_zmult2_eq) lemma sbintrunc_mod2p: "sbintrunc n w = (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n" proof (induction n arbitrary: w) case 0 then show ?case by (auto simp add: odd_iff_mod_2_eq_one) next case (Suc n) moreover have "((bin_rest w + 2 ^ n) mod (2 * 2 ^ n) - 2 ^ n) BIT bin_last w = (w + 2 * 2 ^ n) mod (4 * 2 ^ n) - 2 * 2 ^ n" proof (cases w rule: parity_cases) case even then show ?thesis by (simp add: bin_rest_def Bit_B0_2t mult_mod_right) next case odd then have "2 * (w div 2) = w - 1" using minus_mod_eq_mult_div [of w 2] by simp moreover have "(2 * 2 ^ n + w - 1) mod (2 * 2 * 2 ^ n) + 1 = (2 * 2 ^ n + w) mod (2 * 2 * 2 ^ n)" using odd emep1 [of "2 * 2 ^ n + w - 1" "2 * 2 * 2 ^ n"] by simp ultimately show ?thesis using odd by (simp add: bin_rest_def Bit_B1_2t mult_mod_right) (simp add: algebra_simps) qed ultimately show ?case by simp qed lemma sign_bintr: "bin_sign (bintrunc n w) = 0" by (simp add: bintrunc_mod2p bin_sign_def) lemma bintrunc_n_0 [simp]: "bintrunc n 0 = 0" by (simp add: bintrunc_mod2p) lemma sbintrunc_n_0 [simp]: "sbintrunc n 0 = 0" by (simp add: sbintrunc_mod2p) lemma sbintrunc_n_minus1 [simp]: "sbintrunc n (- 1) = -1" by (induct n) auto lemma bintrunc_Suc_numeral: "bintrunc (Suc n) 1 = 1" "bintrunc (Suc n) (- 1) = bintrunc n (- 1) BIT True" "bintrunc (Suc n) (numeral (Num.Bit0 w)) = bintrunc n (numeral w) BIT False" "bintrunc (Suc n) (numeral (Num.Bit1 w)) = bintrunc n (numeral w) BIT True" "bintrunc (Suc n) (- numeral (Num.Bit0 w)) = bintrunc n (- numeral w) BIT False" "bintrunc (Suc n) (- numeral (Num.Bit1 w)) = bintrunc n (- numeral (w + Num.One)) BIT True" by simp_all 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)) = sbintrunc n (numeral w) BIT False" "sbintrunc (Suc n) (numeral (Num.Bit1 w)) = sbintrunc n (numeral w) BIT True" "sbintrunc (Suc n) (- numeral (Num.Bit0 w)) = sbintrunc n (- numeral w) BIT False" "sbintrunc (Suc n) (- numeral (Num.Bit1 w)) = sbintrunc n (- numeral (w + Num.One)) BIT True" by simp_all lemma bin_sign_lem: "(bin_sign (sbintrunc n bin) = -1) = bin_nth bin n" apply (induct n arbitrary: bin) apply (case_tac bin rule: bin_exhaust, case_tac b, auto) done lemma nth_bintr: "bin_nth (bintrunc m w) n \ n < m \ bin_nth w n" apply (induct n arbitrary: w m) apply (case_tac m, auto)[1] apply (case_tac m, auto)[1] done lemma nth_sbintr: "bin_nth (sbintrunc m w) n = (if n < m then bin_nth w n else bin_nth w m)" apply (induct n arbitrary: w m) apply (case_tac m) apply simp_all apply (case_tac m) apply simp_all done lemma bin_nth_Bit: "bin_nth (w BIT b) n \ n = 0 \ b \ (\m. n = Suc m \ bin_nth w m)" by (cases n) auto lemma bin_nth_Bit0: "bin_nth (numeral (Num.Bit0 w)) n \ (\m. n = Suc m \ bin_nth (numeral w) m)" using bin_nth_Bit [where w="numeral w" and b="False"] by simp lemma bin_nth_Bit1: "bin_nth (numeral (Num.Bit1 w)) n \ n = 0 \ (\m. n = Suc m \ bin_nth (numeral w) m)" using bin_nth_Bit [where w="numeral w" and b="True"] by simp lemma bintrunc_bintrunc_l: "n \ m \ bintrunc m (bintrunc n w) = bintrunc n w" by (rule bin_eqI) (auto simp: nth_bintr) lemma sbintrunc_sbintrunc_l: "n \ m \ sbintrunc m (sbintrunc n w) = sbintrunc n w" by (rule bin_eqI) (auto simp: nth_sbintr) 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 bintrunc_Pls = bintrunc.Suc [where bin="0", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas bintrunc_Min [simp] = bintrunc.Suc [where bin="-1", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas bintrunc_BIT [simp] = bintrunc.Suc [where bin="w BIT b", simplified bin_last_BIT bin_rest_BIT] for w b lemmas bintrunc_Sucs = bintrunc_Pls bintrunc_Min bintrunc_BIT bintrunc_Suc_numeral lemmas sbintrunc_Suc_Pls = sbintrunc.Suc [where bin="0", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Suc_Min = sbintrunc.Suc [where bin="-1", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Suc_BIT [simp] = sbintrunc.Suc [where bin="w BIT b", simplified bin_last_BIT bin_rest_BIT] for w b lemmas sbintrunc_Sucs = sbintrunc_Suc_Pls sbintrunc_Suc_Min sbintrunc_Suc_BIT sbintrunc_Suc_numeral lemmas sbintrunc_Pls = sbintrunc.Z [where bin="0", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Min = sbintrunc.Z [where bin="-1", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_0_BIT_B0 [simp] = sbintrunc.Z [where bin="w BIT False", simplified bin_last_numeral_simps bin_rest_numeral_simps] for w lemmas sbintrunc_0_BIT_B1 [simp] = sbintrunc.Z [where bin="w BIT True", simplified bin_last_BIT bin_rest_numeral_simps] for w lemmas sbintrunc_0_simps = sbintrunc_Pls sbintrunc_Min sbintrunc_0_BIT_B0 sbintrunc_0_BIT_B1 lemmas bintrunc_simps = bintrunc.Z bintrunc_Sucs 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 bintrunc_minus_simps = bintrunc_Sucs [THEN [2] bintrunc_minus [symmetric, THEN trans]] lemmas sbintrunc_minus_simps = sbintrunc_Sucs [THEN [2] sbintrunc_minus [symmetric, THEN trans]] lemmas thobini1 = arg_cong [where f = "\w. w BIT b"] for b lemmas bintrunc_BIT_I = trans [OF bintrunc_BIT thobini1] lemmas bintrunc_Min_I = trans [OF bintrunc_Min thobini1] lemmas bmsts = bintrunc_minus_simps(1-3) [THEN thobini1 [THEN [2] trans]] lemmas bintrunc_Pls_minus_I = bmsts(1) lemmas bintrunc_Min_minus_I = bmsts(2) lemmas bintrunc_BIT_minus_I = bmsts(3) lemma bintrunc_Suc_lem: "bintrunc (Suc n) x = y \ m = Suc n \ bintrunc m x = y" by auto lemmas bintrunc_Suc_Ialts = bintrunc_Min_I [THEN bintrunc_Suc_lem] bintrunc_BIT_I [THEN bintrunc_Suc_lem] lemmas sbintrunc_BIT_I = trans [OF sbintrunc_Suc_BIT thobini1] lemmas sbintrunc_Suc_Is = sbintrunc_Sucs(1-3) [THEN thobini1 [THEN [2] trans]] lemmas sbintrunc_Suc_minus_Is = sbintrunc_minus_simps(1-3) [THEN thobini1 [THEN [2] trans]] 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) (auto simp del: bintrunc.Suc) lemma sbintrunc_bintrunc' [simp]: "0 < n \ sbintrunc (n - 1) (bintrunc n w) = sbintrunc (n - 1) w" by (cases n) (auto simp del: bintrunc.Suc) 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 del: bintrunc.Suc) 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 = bintrunc (pred_numeral k) (bin_rest x) BIT bin_last x" by (simp add: numeral_eq_Suc) lemma sbintrunc_numeral: "sbintrunc (numeral k) x = sbintrunc (pred_numeral k) (bin_rest x) BIT bin_last x" by (simp add: numeral_eq_Suc) lemma bintrunc_numeral_simps [simp]: "bintrunc (numeral k) (numeral (Num.Bit0 w)) = bintrunc (pred_numeral k) (numeral w) BIT False" "bintrunc (numeral k) (numeral (Num.Bit1 w)) = bintrunc (pred_numeral k) (numeral w) BIT True" "bintrunc (numeral k) (- numeral (Num.Bit0 w)) = bintrunc (pred_numeral k) (- numeral w) BIT False" "bintrunc (numeral k) (- numeral (Num.Bit1 w)) = bintrunc (pred_numeral k) (- numeral (w + Num.One)) BIT True" "bintrunc (numeral k) 1 = 1" by (simp_all add: bintrunc_numeral) lemma sbintrunc_numeral_simps [simp]: "sbintrunc (numeral k) (numeral (Num.Bit0 w)) = sbintrunc (pred_numeral k) (numeral w) BIT False" "sbintrunc (numeral k) (numeral (Num.Bit1 w)) = sbintrunc (pred_numeral k) (numeral w) BIT True" "sbintrunc (numeral k) (- numeral (Num.Bit0 w)) = sbintrunc (pred_numeral k) (- numeral w) BIT False" "sbintrunc (numeral k) (- numeral (Num.Bit1 w)) = sbintrunc (pred_numeral k) (- numeral (w + Num.One)) BIT True" "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}" apply (unfold no_bintr_alt1) apply (auto simp add: image_iff) apply (rule exI) apply (rule sym) using int_mod_lem [symmetric, of "2 ^ n"] apply auto done 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}" apply (unfold no_sbintr_alt2) apply (auto simp add: image_iff eq_diff_eq) apply (rule exI) apply (auto intro: int_mod_lem [THEN iffD1, symmetric]) done lemma sb_inc_lem: "a + 2^k < 0 \ a + 2^k + 2^(Suc k) \ (a + 2^k) mod 2^(Suc k)" for a :: int using int_mod_ge' [where n = "2 ^ (Suc k)" and b = "a + 2 ^ k"] by simp lemma sb_inc_lem': "a < - (2^k) \ a + 2^k + 2^(Suc k) \ (a + 2^k) mod 2^(Suc k)" for a :: int by (rule sb_inc_lem) simp lemma sbintrunc_inc: "x < - (2^n) \ x + 2^(Suc n) \ sbintrunc n x" unfolding no_sbintr_alt2 by (drule sb_inc_lem') simp lemma sb_dec_lem: "0 \ - (2 ^ k) + a \ (a + 2 ^ k) mod (2 * 2 ^ k) \ - (2 ^ k) + a" for a :: int using int_mod_le'[where n = "2 ^ (Suc k)" and b = "a + 2 ^ k"] by simp lemma sb_dec_lem': "2 ^ k \ a \ (a + 2 ^ k) mod (2 * 2 ^ k) \ - (2 ^ k) + a" for a :: int by (rule sb_dec_lem) simp lemma sbintrunc_dec: "x \ (2 ^ n) \ x - 2 ^ (Suc n) >= sbintrunc n x" unfolding no_sbintr_alt2 by (drule sb_dec_lem') simp 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: bintrunc_mod2p m1mod2k) 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 (induct n arbitrary: bin) auto 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 lemma bin_rest_strunc: "bin_rest (sbintrunc (Suc n) bin) = sbintrunc n (bin_rest bin)" by (induct n arbitrary: bin) auto lemma bintrunc_rest [simp]: "bintrunc n (bin_rest (bintrunc n bin)) = bin_rest (bintrunc n bin)" apply (induct n arbitrary: bin) apply simp apply (case_tac bin rule: bin_exhaust) apply (auto simp: bintrunc_bintrunc_l) done lemma sbintrunc_rest [simp]: "sbintrunc n (bin_rest (sbintrunc n bin)) = bin_rest (sbintrunc n bin)" apply (induct n arbitrary: bin) apply simp apply (case_tac bin rule: bin_exhaust) apply (auto simp: bintrunc_bintrunc_l split: bool.splits) done 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\ primrec bin_split :: "nat \ int \ int \ int" where Z: "bin_split 0 w = (w, 0)" | Suc: "bin_split (Suc n) w = (let (w1, w2) = bin_split n (bin_rest w) in (w1, w2 BIT bin_last w))" lemma [code]: "bin_split (Suc n) w = (let (w1, w2) = bin_split n (bin_rest w) in (w1, w2 BIT bin_last w))" "bin_split 0 w = (w, 0)" by simp_all primrec bin_cat :: "int \ nat \ int \ int" where Z: "bin_cat w 0 v = w" | Suc: "bin_cat w (Suc n) v = bin_cat w n (bin_rest v) BIT bin_last v" lemma bin_sign_cat: "bin_sign (bin_cat x n y) = bin_sign x" by (induct n arbitrary: y) auto lemma bin_cat_Suc_Bit: "bin_cat w (Suc n) (v BIT b) = bin_cat w n v BIT b" by auto lemma bin_cat_assoc: "bin_cat (bin_cat x m y) n z = bin_cat x (m + n) (bin_cat y n z)" by (induct n arbitrary: z) auto 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" apply (induct n arbitrary: z m) apply clarsimp apply (case_tac m, auto) done 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))" apply (induct k arbitrary: n y) apply clarsimp apply (case_tac n, auto) done 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))" apply (induct n arbitrary: b c) apply clarsimp apply (clarsimp simp: Let_def split: prod.split_asm) apply (case_tac k) apply auto done lemma bin_cat_zero [simp]: "bin_cat 0 n w = bintrunc n w" by (induct n arbitrary: w) auto lemma bintr_cat1: "bintrunc (k + n) (bin_cat a n b) = bin_cat (bintrunc k a) n b" by (induct n arbitrary: b) auto 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 (induct n arbitrary: b) auto lemma split_bintrunc: "bin_split n c = (a, b) \ b = bintrunc n c" by (induct n arbitrary: b c) (auto simp: Let_def split: prod.split_asm) lemma bin_cat_split: "bin_split n w = (u, v) \ w = bin_cat u n v" by (induct n arbitrary: v w) (auto simp: Let_def split: prod.split_asm) lemma bin_split_cat: "bin_split n (bin_cat v n w) = (v, bintrunc n w)" by (induct n arbitrary: w) auto lemma bin_split_zero [simp]: "bin_split n 0 = (0, 0)" by (induct n) auto lemma bin_split_minus1 [simp]: "bin_split n (- 1) = (- 1, bintrunc n (- 1))" by (induct n) auto 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 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 split: prod.split_asm) done lemma bin_cat_num: "bin_cat a n b = a * 2 ^ n + bintrunc n b" apply (induct n arbitrary: b) apply clarsimp apply (simp add: Bit_def) done lemma bin_split_num: "bin_split n b = (b div 2 ^ n, b mod 2 ^ n)" apply (induct n arbitrary: b) apply simp apply (simp add: bin_rest_def zdiv_zmult2_eq) apply (case_tac b rule: bin_exhaust) apply simp apply (simp add: Bit_def mod_mult_mult1 pos_zmod_mult_2 add.commute) done 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 lemmas bin_split_minus_simp = bin_split.Suc [THEN [2] bin_split_minus [symmetric, THEN trans]] 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, w2 BIT bin_last w))" by (simp only: bin_split_minus_simp) 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 (drule split_bintrunc) 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 clarify 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 (drule bin_nth_split, erule conjE, erule allE, erule trans, simp add: 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 (drule bin_split_trunc) apply (drule sym [THEN trans], assumption) apply (subst rsplit_aux_alts(1)) apply (subst rsplit_aux_alts(2)) apply clarsimp unfolding bin_rsplit_def bin_rsplitl_def apply simp 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) unfolding bin_split_cat apply simp 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" \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) (fst (bin_split n w)) (snd (bin_split n w) # cs))" by auto 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 = bin_rest w BIT b" | Suc: "bin_sc (Suc n) b w = bin_sc n b (bin_rest w) BIT bin_last w" lemma bin_nth_sc [simp]: "bin_nth (bin_sc n b w) n \ b" by (induct n arbitrary: w) auto lemma bin_sc_sc_same [simp]: "bin_sc n c (bin_sc n b w) = bin_sc n c w" by (induct n arbitrary: w) auto 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)" by (induct n arbitrary: w m) (case_tac [!] m, auto) lemma bin_sc_nth [simp]: "bin_sc n (bin_nth w n) w = w" by (induct n arbitrary: w) auto lemma bin_sign_sc [simp]: "bin_sign (bin_sc n b w) = bin_sign w" by (induct n arbitrary: w) auto lemma bin_sc_bintr [simp]: "bintrunc m (bin_sc n x (bintrunc m (w))) = bintrunc m (bin_sc n x w)" apply (induct n arbitrary: w m) apply (case_tac [!] w rule: bin_exhaust) apply (case_tac [!] m, auto) done lemma bin_clr_le: "bin_sc n False w \ w" apply (induct n arbitrary: w) apply (case_tac [!] w rule: bin_exhaust) apply (auto simp: le_Bits) done lemma bin_set_ge: "bin_sc n True w \ w" apply (induct n arbitrary: w) apply (case_tac [!] w rule: bin_exhaust) apply (auto simp: le_Bits) done lemma bintr_bin_clr_le: "bintrunc n (bin_sc m False w) \ bintrunc n w" apply (induct n arbitrary: w m) apply simp apply (case_tac w rule: bin_exhaust) apply (case_tac m) apply (auto simp: le_Bits) done lemma bintr_bin_set_ge: "bintrunc n (bin_sc m True w) \ bintrunc n w" apply (induct n arbitrary: w m) apply simp apply (case_tac w rule: bin_exhaust) apply (case_tac m) apply (auto simp: le_Bits) done 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 = bin_sc (pred_numeral k) b (bin_rest w) BIT bin_last w" by (simp add: numeral_eq_Suc) instantiation int :: bit_operations begin definition int_not_def: "NOT = (\x::int. - x - 1)" function bitAND_int where "bitAND_int x y = (if x = 0 then 0 else if x = -1 then y else (bin_rest x AND bin_rest y) BIT (bin_last x \ bin_last y))" by pat_completeness simp termination by (relation "measure (nat \ abs \ fst)", simp_all add: bin_rest_def) declare bitAND_int.simps [simp del] definition int_or_def: "(OR) = (\x y::int. NOT (NOT x AND NOT y))" definition int_xor_def: "(XOR) = (\x y::int. (x AND NOT y) OR (NOT x AND y))" definition [iff]: "i !! n \ bin_nth i n" definition "lsb i = i !! 0" for i :: int definition "set_bit i n b = bin_sc n b i" definition "shiftl x n = x * 2 ^ n" for x :: int definition "shiftr x n = x div 2 ^ n" for x :: int definition "msb x \ x < 0" for x :: int instance .. end subsubsection \Basic simplification rules\ lemma int_not_BIT [simp]: "NOT (w BIT b) = (NOT w) BIT (\ b)" by (cases b) (simp_all add: int_not_def Bit_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)" unfolding int_not_def by simp_all lemma int_not_not [simp]: "NOT (NOT x) = x" for x :: int unfolding int_not_def by simp lemma int_and_0 [simp]: "0 AND x = 0" for x :: int by (simp add: bitAND_int.simps) lemma int_and_m1 [simp]: "-1 AND x = x" for x :: int by (simp add: bitAND_int.simps) lemma int_and_Bits [simp]: "(x BIT b) AND (y BIT c) = (x AND y) BIT (b \ c)" by (subst bitAND_int.simps) (simp add: Bit_eq_0_iff Bit_eq_m1_iff) lemma int_or_zero [simp]: "0 OR x = x" for x :: int by (simp add: int_or_def) lemma int_or_minus1 [simp]: "-1 OR x = -1" for x :: int by (simp add: int_or_def) lemma int_or_Bits [simp]: "(x BIT b) OR (y BIT c) = (x OR y) BIT (b \ c)" by (simp add: int_or_def) lemma int_xor_zero [simp]: "0 XOR x = x" for x :: int by (simp add: int_xor_def) lemma int_xor_Bits [simp]: "(x BIT b) XOR (y BIT c) = (x XOR y) BIT ((b \ c) \ \ (b \ c))" unfolding int_xor_def by auto subsubsection \Binary destructors\ lemma bin_rest_NOT [simp]: "bin_rest (NOT x) = NOT (bin_rest x)" by (cases x rule: bin_exhaust) simp lemma bin_last_NOT [simp]: "bin_last (NOT x) \ \ bin_last x" by (cases x rule: bin_exhaust) simp lemma bin_rest_AND [simp]: "bin_rest (x AND y) = bin_rest x AND bin_rest y" by (cases x rule: bin_exhaust, cases y rule: bin_exhaust) simp lemma bin_last_AND [simp]: "bin_last (x AND y) \ bin_last x \ bin_last y" by (cases x rule: bin_exhaust, cases y rule: bin_exhaust) simp lemma bin_rest_OR [simp]: "bin_rest (x OR y) = bin_rest x OR bin_rest y" by (cases x rule: bin_exhaust, cases y rule: bin_exhaust) simp lemma bin_last_OR [simp]: "bin_last (x OR y) \ bin_last x \ bin_last y" by (cases x rule: bin_exhaust, cases y rule: bin_exhaust) simp lemma bin_rest_XOR [simp]: "bin_rest (x XOR y) = bin_rest x XOR bin_rest y" by (cases x rule: bin_exhaust, cases y rule: bin_exhaust) simp lemma bin_last_XOR [simp]: "bin_last (x XOR y) \ (bin_last x \ bin_last y) \ \ (bin_last x \ bin_last y)" by (cases x rule: bin_exhaust, cases y rule: bin_exhaust) simp 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 (induct n) auto subsubsection \Derived properties\ lemma int_xor_minus1 [simp]: "-1 XOR x = NOT x" for x :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_xor_extra_simps [simp]: "w XOR 0 = w" "w XOR -1 = NOT w" for w :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_or_extra_simps [simp]: "w OR 0 = w" "w OR -1 = -1" for w :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_and_extra_simps [simp]: "w AND 0 = 0" "w AND -1 = w" for w :: int by (auto simp add: bin_eq_iff bin_nth_ops) 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 (auto simp add: bin_eq_iff bin_nth_ops) lemma bin_ops_same [simp]: "x AND x = x" "x OR x = x" "x XOR x = 0" for x :: int by (auto simp add: bin_eq_iff bin_nth_ops) 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 flip: BIT_bin_simps) 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) = (numeral x AND numeral y) BIT False" "numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (numeral x AND numeral y) BIT False" "numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (numeral x AND numeral y) BIT False" "numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = (numeral x AND numeral y) BIT True" "numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (numeral x AND - numeral y) BIT False" "numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (numeral x AND - numeral (y + Num.One)) BIT False" "numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (numeral x AND - numeral y) BIT False" "numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = (numeral x AND - numeral (y + Num.One)) BIT True" "- numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (- numeral x AND numeral y) BIT False" "- numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (- numeral x AND numeral y) BIT False" "- numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (- numeral (x + Num.One) AND numeral y) BIT False" "- numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = (- numeral (x + Num.One) AND numeral y) BIT True" "- numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (- numeral x AND - numeral y) BIT False" "- numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (- numeral x AND - numeral (y + Num.One)) BIT False" "- numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (- numeral (x + Num.One) AND - numeral y) BIT False" "- numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = (- numeral (x + Num.One) AND - numeral (y + Num.One)) BIT True" "(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) = (numeral x OR numeral y) BIT False" "numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = (numeral x OR numeral y) BIT True" "numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = (numeral x OR numeral y) BIT True" "numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = (numeral x OR numeral y) BIT True" "numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (numeral x OR - numeral y) BIT False" "numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = (numeral x OR - numeral (y + Num.One)) BIT True" "numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = (numeral x OR - numeral y) BIT True" "numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = (numeral x OR - numeral (y + Num.One)) BIT True" "- numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (- numeral x OR numeral y) BIT False" "- numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = (- numeral x OR numeral y) BIT True" "- numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = (- numeral (x + Num.One) OR numeral y) BIT True" "- numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = (- numeral (x + Num.One) OR numeral y) BIT True" "- numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (- numeral x OR - numeral y) BIT False" "- numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = (- numeral x OR - numeral (y + Num.One)) BIT True" "- numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = (- numeral (x + Num.One) OR - numeral y) BIT True" "- numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = (- numeral (x + Num.One) OR - numeral (y + Num.One)) BIT True" "(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) = (numeral x XOR numeral y) BIT False" "numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = (numeral x XOR numeral y) BIT True" "numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = (numeral x XOR numeral y) BIT True" "numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (numeral x XOR numeral y) BIT False" "numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (numeral x XOR - numeral y) BIT False" "numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = (numeral x XOR - numeral (y + Num.One)) BIT True" "numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = (numeral x XOR - numeral y) BIT True" "numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (numeral x XOR - numeral (y + Num.One)) BIT False" "- numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (- numeral x XOR numeral y) BIT False" "- numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = (- numeral x XOR numeral y) BIT True" "- numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = (- numeral (x + Num.One) XOR numeral y) BIT True" "- numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (- numeral (x + Num.One) XOR numeral y) BIT False" "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (- numeral x XOR - numeral y) BIT False" "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = (- numeral x XOR - numeral (y + Num.One)) BIT True" "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = (- numeral (x + Num.One) XOR - numeral y) BIT True" "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (- numeral (x + Num.One) XOR - numeral (y + Num.One)) BIT False" "(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 [rule_format]: "\y::int. (x AND y) + (x OR y) = x + y" apply (induct x rule: bin_induct) apply clarsimp apply clarsimp apply clarsimp apply (case_tac y rule: bin_exhaust) apply clarsimp apply (unfold Bit_def) apply clarsimp apply (erule_tac x = "x" in allE) apply simp done lemma le_int_or: "bin_sign y = 0 \ x \ x OR y" for x y :: int apply (induct y arbitrary: x rule: bin_induct) apply clarsimp apply clarsimp apply (case_tac x rule: bin_exhaust) apply (case_tac b) apply (case_tac [!] bit) apply (auto simp: le_Bits) done 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)" apply (induct x rule: bin_induct) apply clarsimp apply clarsimp apply (case_tac bit, auto) done lemma mod_BIT: "bin BIT bit mod 2 ^ Suc n = (bin mod 2 ^ n) BIT bit" for bit proof - have "2 * (bin mod 2 ^ n) + 1 = (2 * bin mod 2 ^ Suc n) + 1" by (simp add: mod_mult_mult1) also have "\ = ((2 * bin mod 2 ^ Suc n) + 1) mod 2 ^ Suc n" by (simp add: ac_simps pos_zmod_mult_2) also have "\ = (2 * bin + 1) mod 2 ^ Suc n" by (simp only: mod_simps) finally show ?thesis by (auto simp add: Bit_def) qed lemma AND_mod: "x AND 2 ^ n - 1 = x mod 2 ^ n" for x :: int proof (induct x arbitrary: n rule: bin_induct) case 1 then show ?case by simp next case 2 then show ?case by (simp, simp add: m1mod2k) next case (3 bin bit) show ?case proof (cases n) case 0 then show ?thesis by simp next case (Suc m) with 3 show ?thesis by (simp only: power_BIT mod_BIT int_and_Bits) simp qed qed subsubsection \Comparison\ lemma AND_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" shows "0 \ x AND y" using assms proof (induct x arbitrary: y rule: bin_induct) case 1 then show ?case by simp next case 2 then show ?case by (simp only: Min_def) next case (3 bin bit) show ?case proof (cases y rule: bin_exhaust) case (1 bin' bit') from 3 have "0 \ bin" by (cases bit) (simp_all add: Bit_def) then have "0 \ bin AND bin'" by (rule 3) with 1 show ?thesis by simp qed qed lemma OR_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" "0 \ y" shows "0 \ x OR y" using assms proof (induct x arbitrary: y rule: bin_induct) case (3 bin bit) show ?case proof (cases y rule: bin_exhaust) case (1 bin' bit') from 3 have "0 \ bin" by (cases bit) (simp_all add: Bit_def) moreover from 1 3 have "0 \ bin'" by (cases bit') (simp_all add: Bit_def) ultimately have "0 \ bin OR bin'" by (rule 3) with 1 show ?thesis by simp qed qed simp_all lemma XOR_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" "0 \ y" shows "0 \ x XOR y" using assms proof (induct x arbitrary: y rule: bin_induct) case (3 bin bit) show ?case proof (cases y rule: bin_exhaust) case (1 bin' bit') from 3 have "0 \ bin" by (cases bit) (simp_all add: Bit_def) moreover from 1 3 have "0 \ bin'" by (cases bit') (simp_all add: Bit_def) ultimately have "0 \ bin XOR bin'" by (rule 3) with 1 show ?thesis by simp qed next case 2 then show ?case by (simp only: Min_def) qed simp lemma AND_upper1 [simp]: \<^marker>\contributor \Stefan Berghofer\\ fixes x y :: int assumes "0 \ x" shows "x AND y \ x" using assms proof (induct x arbitrary: y rule: bin_induct) case (3 bin bit) show ?case proof (cases y rule: bin_exhaust) case (1 bin' bit') from 3 have "0 \ bin" by (cases bit) (simp_all add: Bit_def) then have "bin AND bin' \ bin" by (rule 3) with 1 show ?thesis by simp (simp add: Bit_def) qed next case 2 then show ?case by (simp only: Min_def) qed simp 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 proof (induct y arbitrary: x rule: bin_induct) case 1 then show ?case by simp next case 2 then show ?case by (simp only: Min_def) next case (3 bin bit) show ?case proof (cases x rule: bin_exhaust) case (1 bin' bit') from 3 have "0 \ bin" by (cases bit) (simp_all add: Bit_def) then have "bin' AND bin \ bin" by (rule 3) with 1 show ?thesis by simp (simp add: Bit_def) qed qed 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 (induct x arbitrary: y n rule: bin_induct) case (3 bin bit) show ?case proof (cases y rule: bin_exhaust) case (1 bin' bit') show ?thesis proof (cases n) case 0 with 3 have "bin BIT bit = 0" by (simp add: Bit_def) then have "bin = 0" and "\ bit" by (auto simp add: Bit_def split: if_splits) arith then show ?thesis using 0 1 \y < 2 ^ n\ by simp next case (Suc m) from 3 have "0 \ bin" by (cases bit) (simp_all add: Bit_def) moreover from 3 Suc have "bin < 2 ^ m" by (cases bit) (simp_all add: Bit_def) moreover from 1 3 Suc have "bin' < 2 ^ m" by (cases bit') (simp_all add: Bit_def) ultimately have "bin OR bin' < 2 ^ m" by (rule 3) with 1 Suc show ?thesis by simp (simp add: Bit_def) qed qed qed simp_all 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 (induct x arbitrary: y n rule: bin_induct) case 1 then show ?case by simp next case 2 then show ?case by (simp only: Min_def) next case (3 bin bit) show ?case proof (cases y rule: bin_exhaust) case (1 bin' bit') show ?thesis proof (cases n) case 0 with 3 have "bin BIT bit = 0" by (simp add: Bit_def) then have "bin = 0" and "\ bit" by (auto simp add: Bit_def split: if_splits) arith then show ?thesis using 0 1 \y < 2 ^ n\ by simp next case (Suc m) from 3 have "0 \ bin" by (cases bit) (simp_all add: Bit_def) moreover from 3 Suc have "bin < 2 ^ m" by (cases bit) (simp_all add: Bit_def) moreover from 1 3 Suc have "bin' < 2 ^ m" by (cases bit') (simp_all add: Bit_def) ultimately have "bin XOR bin' < 2 ^ m" by (rule 3) with 1 Suc show ?thesis by simp (simp add: Bit_def) qed qed 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(metis int_and_comm bbw_ao_dist) lemmas int_and_ac = bbw_lcs(1) int_and_comm int_and_assoc lemma int_nand_same [simp]: fixes x :: int shows "x AND NOT x = 0" by(induct x y\"NOT x" rule: bitAND_int.induct)(subst bitAND_int.simps, clarsimp) lemma int_nand_same_middle: fixes x :: int shows "x AND y AND NOT x = 0" by (metis bbw_lcs(1) int_and_0 int_nand_same) lemma and_xor_dist: fixes x :: int shows "x AND (y XOR z) = (x AND y) XOR (x AND z)" by(simp add: int_xor_def bbw_ao_dist2 bbw_not_dist int_and_ac int_nand_same_middle) lemma int_and_lt0 [simp]: fixes x y :: int shows "x AND y < 0 \ x < 0 \ y < 0" by(induct x y rule: bitAND_int.induct)(subst bitAND_int.simps, simp) lemma int_and_ge0 [simp]: fixes x y :: int shows "x AND y \ 0 \ x \ 0 \ y \ 0" by (metis int_and_lt0 linorder_not_less) lemma int_and_1: fixes x :: int shows "x AND 1 = x mod 2" by(subst bitAND_int.simps)(simp add: Bit_def bin_last_def zmod_minus1) lemma int_1_and: fixes x :: int shows "1 AND x = x mod 2" by(subst int_and_comm)(simp add: int_and_1) lemma int_or_lt0 [simp]: fixes x y :: int shows "x OR y < 0 \ x < 0 \ y < 0" by(simp add: int_or_def) lemma int_xor_lt0 [simp]: fixes x y :: int shows "x XOR y < 0 \ ((x < 0) \ (y < 0))" by(auto simp add: int_xor_def) lemma int_xor_ge0 [simp]: fixes x y :: int shows "x XOR y \ 0 \ ((x \ 0) \ (y \ 0))" by (metis int_xor_lt0 linorder_not_le) lemma even_conv_AND: \even i \ i AND 1 = 0\ for i :: int proof - obtain x b where \i = x BIT b\ by (cases i rule: bin_exhaust) then have "i AND 1 = 0 BIT b" by (simp add: BIT_special_simps(2) [symmetric] del: BIT_special_simps(2)) then show ?thesis using \i = x BIT b\ by (cases b) simp_all qed lemma bin_last_conv_AND: "bin_last i \ i AND 1 \ 0" by (simp add: even_conv_AND) lemma bitval_bin_last: "of_bool (bin_last i) = i AND 1" proof - obtain x b where "i = x BIT b" by(cases i rule: bin_exhaust) hence "i AND 1 = 0 BIT b" by(simp add: BIT_special_simps(2)[symmetric] del: BIT_special_simps(2)) thus ?thesis by(cases b)(simp_all add: bin_last_conv_AND) qed lemma bin_sign_and: "bin_sign (i AND j) = - (bin_sign i * bin_sign j)" by(simp add: bin_sign_def) lemma minus_BIT_0: fixes x y :: int shows "x BIT b - y BIT False = (x - y) BIT b" by(simp add: Bit_def) lemma int_not_neg_numeral: "NOT (- numeral n) = (Num.sub n num.One :: int)" by(simp add: int_not_def) lemma 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_lsb_BIT [simp]: fixes x :: int shows "lsb (x BIT b) \ b" by(simp add: lsb_int_def) lemma bin_last_conv_lsb: "bin_last = lsb" by(clarsimp simp add: lsb_int_def fun_eq_iff) lemma int_lsb_numeral [simp]: "lsb (0 :: int) = False" "lsb (1 :: int) = True" "lsb (Numeral1 :: int) = True" "lsb (- 1 :: int) = True" "lsb (- Numeral1 :: int) = True" "lsb (numeral (num.Bit0 w) :: int) = False" "lsb (numeral (num.Bit1 w) :: int) = True" "lsb (- numeral (num.Bit0 w) :: int) = False" "lsb (- numeral (num.Bit1 w) :: int) = True" by (simp_all add: lsb_int_def) lemma int_set_bit_0 [simp]: fixes x :: int shows "set_bit x 0 b = bin_rest x BIT b" by(auto simp add: set_bit_int_def intro: bin_rl_eqI) lemma int_set_bit_Suc: fixes x :: int shows "set_bit x (Suc n) b = set_bit (bin_rest x) n b BIT bin_last x" by(auto simp add: set_bit_int_def twice_conv_BIT intro: bin_rl_eqI) lemma bin_last_set_bit: "bin_last (set_bit x n b) = (if n > 0 then bin_last x else b)" by(cases n)(simp_all add: int_set_bit_Suc) lemma bin_rest_set_bit: "bin_rest (set_bit x n b) = (if n > 0 then set_bit (bin_rest x) (n - 1) b else bin_rest x)" by(cases n)(simp_all add: int_set_bit_Suc) lemma int_set_bit_numeral: fixes x :: int shows "set_bit x (numeral w) b = set_bit (bin_rest x) (pred_numeral w) b BIT bin_last x" by(simp add: set_bit_int_def) lemmas int_set_bit_numerals [simp] = int_set_bit_numeral[where x="numeral w'"] int_set_bit_numeral[where x="- numeral w'"] int_set_bit_numeral[where x="Numeral1"] int_set_bit_numeral[where x="1"] int_set_bit_numeral[where x="0"] int_set_bit_Suc[where x="numeral w'"] int_set_bit_Suc[where x="- numeral w'"] int_set_bit_Suc[where x="Numeral1"] int_set_bit_Suc[where x="1"] int_set_bit_Suc[where x="0"] for w' lemma int_shiftl_BIT: fixes x :: int shows int_shiftl0 [simp]: "x << 0 = x" and int_shiftl_Suc [simp]: "x << Suc n = (x << n) BIT False" by(auto simp add: shiftl_int_def Bit_def) lemma int_0_shiftl [simp]: "0 << n = (0 :: int)" by(induct n) simp_all lemma bin_last_shiftl: "bin_last (x << n) \ n = 0 \ bin_last x" by(cases n)(simp_all) lemma bin_rest_shiftl: "bin_rest (x << n) = (if n > 0 then x << (n - 1) else bin_rest x)" by(cases n)(simp_all) lemma bin_nth_shiftl [simp]: "bin_nth (x << n) m \ n \ m \ bin_nth x (m - n)" proof(induct n arbitrary: x m) case (Suc n) thus ?case by(cases m) simp_all qed simp lemma int_shiftr_BIT [simp]: fixes x :: int shows int_shiftr0: "x >> 0 = x" and int_shiftr_Suc: "x BIT b >> Suc n = x >> n" proof - show "x >> 0 = x" by (simp add: shiftr_int_def) show "x BIT b >> Suc n = x >> n" by (cases b) (simp_all add: shiftr_int_def Bit_def add.commute pos_zdiv_mult_2) qed lemma bin_last_shiftr: "bin_last (x >> n) \ x !! n" proof(induct n arbitrary: x) case 0 thus ?case by simp next case (Suc n) thus ?case by(cases x rule: bin_exhaust) simp qed lemma bin_rest_shiftr [simp]: "bin_rest (x >> n) = x >> Suc n" proof(induct n arbitrary: x) case 0 thus ?case by(cases x rule: bin_exhaust) auto next case (Suc n) thus ?case by(cases x rule: bin_exhaust) auto qed lemma bin_nth_shiftr [simp]: "bin_nth (x >> n) m = bin_nth x (n + m)" proof(induct n arbitrary: x m) case (Suc n) thus ?case by(cases x rule: bin_exhaust) simp_all qed simp lemma bin_nth_conv_AND: fixes x :: int shows "bin_nth x n \ x AND (1 << n) \ 0" proof(induct n arbitrary: x) case 0 thus ?case by(simp add: int_and_1 bin_last_def) next case (Suc n) thus ?case by(cases x rule: bin_exhaust)(simp_all add: bin_nth_ops Bit_eq_0_iff) qed lemma int_shiftl_numeral [simp]: "(numeral w :: int) << numeral w' = numeral (num.Bit0 w) << pred_numeral w'" "(- numeral w :: int) << numeral w' = - numeral (num.Bit0 w) << pred_numeral w'" by(simp_all add: numeral_eq_Suc Bit_def shiftl_int_def) (metis add_One mult_inc semiring_norm(11) semiring_norm(13) semiring_norm(2) semiring_norm(6) semiring_norm(87))+ lemma int_shiftl_One_numeral [simp]: "(1 :: int) << numeral w = 2 << pred_numeral w" by(metis int_shiftl_numeral numeral_One) lemma shiftl_ge_0 [simp]: fixes i :: int shows "i << n \ 0 \ i \ 0" by(induct n) simp_all lemma shiftl_lt_0 [simp]: fixes i :: int shows "i << n < 0 \ i < 0" by (metis not_le shiftl_ge_0) lemma int_shiftl_test_bit: "(n << i :: int) !! m \ m \ i \ n !! (m - i)" proof(induction i) case (Suc n) thus ?case by(cases m) simp_all qed simp lemma int_0shiftr [simp]: "(0 :: int) >> x = 0" by(simp add: shiftr_int_def) lemma int_minus1_shiftr [simp]: "(-1 :: int) >> x = -1" by(simp add: shiftr_int_def div_eq_minus1) lemma int_shiftr_ge_0 [simp]: fixes i :: int shows "i >> n \ 0 \ i \ 0" proof(induct n arbitrary: i) case (Suc n) thus ?case by(cases i rule: bin_exhaust) simp_all qed simp lemma int_shiftr_lt_0 [simp]: fixes i :: int shows "i >> n < 0 \ i < 0" by (metis int_shiftr_ge_0 not_less) lemma int_shiftr_numeral [simp]: "(1 :: int) >> numeral w' = 0" "(numeral num.One :: int) >> numeral w' = 0" "(numeral (num.Bit0 w) :: int) >> numeral w' = numeral w >> pred_numeral w'" "(numeral (num.Bit1 w) :: int) >> numeral w' = numeral w >> pred_numeral w'" "(- numeral (num.Bit0 w) :: int) >> numeral w' = - numeral w >> pred_numeral w'" "(- numeral (num.Bit1 w) :: int) >> numeral w' = - numeral (Num.inc w) >> pred_numeral w'" by (simp_all only: numeral_One expand_BIT numeral_eq_Suc int_shiftr_Suc BIT_special_simps(2)[symmetric] int_0shiftr add_One uminus_Bit_eq) (simp_all add: add_One) lemma int_shiftr_numeral_Suc0 [simp]: "(1 :: int) >> Suc 0 = 0" "(numeral num.One :: int) >> Suc 0 = 0" "(numeral (num.Bit0 w) :: int) >> Suc 0 = numeral w" "(numeral (num.Bit1 w) :: int) >> Suc 0 = numeral w" "(- numeral (num.Bit0 w) :: int) >> Suc 0 = - numeral w" "(- numeral (num.Bit1 w) :: int) >> Suc 0 = - numeral (Num.inc w)" by(simp_all only: One_nat_def[symmetric] numeral_One[symmetric] int_shiftr_numeral pred_numeral_simps int_shiftr0) lemma bin_nth_minus_p2: assumes sign: "bin_sign x = 0" and y: "y = 1 << n" and m: "m < n" and x: "x < y" shows "bin_nth (x - y) m = bin_nth x m" using sign m x unfolding y proof(induction m arbitrary: x y n) case 0 thus ?case by (simp add: bin_last_def shiftl_int_def) next case (Suc m) from \Suc m < n\ obtain n' where [simp]: "n = Suc n'" by(cases n) auto obtain x' b where [simp]: "x = x' BIT b" by(cases x rule: bin_exhaust) from \bin_sign x = 0\ have "bin_sign x' = 0" by simp moreover from \x < 1 << n\ have "x' < 1 << n'" by(cases b)(simp_all add: Bit_def shiftl_int_def) moreover have "(2 * x' + of_bool b - 2 * 2 ^ n') div 2 = x' + (- (2 ^ n') + of_bool b div 2)" by(simp only: add_diff_eq[symmetric] add.commute div_mult_self2[OF zero_neq_numeral[symmetric]]) ultimately show ?case using Suc.IH[of x' n'] Suc.prems by(cases b)(simp_all add: Bit_def bin_rest_def shiftl_int_def) qed lemma bin_clr_conv_NAND: "bin_sc n False i = i AND NOT (1 << n)" by (induct n arbitrary: i) (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)+ lemma msb_conv_bin_sign: "msb x \ bin_sign x = -1" by(simp add: bin_sign_def not_le msb_int_def) lemma msb_BIT [simp]: "msb (x BIT b) = msb x" by(simp add: msb_int_def) lemma msb_bin_rest [simp]: "msb (bin_rest x) = msb x" by(simp add: msb_int_def) lemma int_msb_and [simp]: "msb ((x :: int) AND y) \ msb x \ msb y" by(simp add: msb_int_def) lemma int_msb_or [simp]: "msb ((x :: int) OR y) \ msb x \ msb y" by(simp add: msb_int_def) lemma int_msb_xor [simp]: "msb ((x :: int) XOR y) \ msb x \ msb y" by(simp add: msb_int_def) lemma int_msb_not [simp]: "msb (NOT (x :: int)) \ \ msb x" by(simp add: msb_int_def not_less) lemma msb_shiftl [simp]: "msb ((x :: int) << n) \ msb x" by(simp add: msb_int_def) lemma msb_shiftr [simp]: "msb ((x :: int) >> r) \ msb x" by(simp add: msb_int_def) lemma msb_bin_sc [simp]: "msb (bin_sc n b x) \ msb x" by(simp add: msb_conv_bin_sign) lemma msb_set_bit [simp]: "msb (set_bit (x :: int) n b) \ msb x" by(simp add: msb_conv_bin_sign set_bit_int_def) lemma msb_0 [simp]: "msb (0 :: int) = False" by(simp add: msb_int_def) lemma msb_1 [simp]: "msb (1 :: int) = False" by(simp add: msb_int_def) lemma msb_numeral [simp]: "msb (numeral n :: int) = False" "msb (- numeral n :: int) = True" by(simp_all add: msb_int_def) 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) 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 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 (induct bs arbitrary: w) auto 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)" apply (induct n arbitrary: w bs) apply clarsimp apply (cases w rule: bin_exhaust) apply simp done 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 (induct bl arbitrary: w) apply clarsimp apply clarsimp apply (cut_tac x=n and y="size bl" in linorder_less_linear) apply (erule disjE, simp add: nth_append)+ apply auto 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 clarsimp apply (case_tac m, clarsimp) apply (clarsimp simp: bin_to_bl_def) apply (simp add: bin_to_bl_aux_alt) done lemma nth_bin_to_bl_aux: "n < m + length bl \ (bin_to_bl_aux m w bl) ! n = (if n < m then bin_nth w (m - 1 - n) else bl ! (n - m))" apply (induct m arbitrary: w n bl) apply clarsimp apply clarsimp apply (case_tac w rule: bin_exhaust) apply simp 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)" apply (induct bs arbitrary: w) apply clarsimp apply clarsimp apply (drule meta_spec, erule xtrans(8) [rotated], simp add: Bit_def)+ done 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)" apply (induct bs arbitrary: w) apply clarsimp apply clarsimp apply (drule meta_spec, erule order_trans [rotated], simp add: Bit_B0_2t Bit_B1_2t algebra_simps)+ apply (simp add: Bit_def) done 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 w rule: bin_exhaust) 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 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 (induct n arbitrary: m bin bs, clarsimp) apply clarsimp apply (case_tac bin rule: bin_exhaust) apply (case_tac "m \ n", simp) apply (case_tac "m - n", simp) apply simp apply (rule_tac f = "\nat. drop nat bs" in arg_cong) apply simp 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) done 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)" by (auto elim: bin_split_take) 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 nth_rev 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_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_bin_bl_rtf takefill_alt rev_take) lemma bl_to_bin_aux_cat: "\nv v. bl_to_bin_aux bs (bin_cat w nv v) = bin_cat w (nv + length bs) (bl_to_bin_aux bs v)" by (induct bs) (simp, simp add: bin_cat_Suc_Bit [symmetric] del: bin_cat.simps) lemma bin_to_bl_aux_cat: "\w bs. 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 (induct nw) auto 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 add: Bit_B0_2t Bit_B1_2t) 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))" apply (unfold bin_to_bl_def) apply (induct n arbitrary: bin) apply simp apply clarsimp apply (case_tac bin rule: bin_exhaust) apply (case_tac b) apply (clarsimp simp: bin_to_bl_aux_alt)+ done 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 (induct n arbitrary: bin) apply simp apply clarsimp apply (case_tac bin rule: bin_exhaust) apply (case_tac b) apply (clarsimp simp: bin_to_bl_aux_alt)+ 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) (x BIT b))" 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 apply (unfold bin_to_bl_def) 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: bin_to_bl_aux_alt Let_def) apply (auto simp: rbbl_Cons rbl_mult_Suc rbl_add) 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 (induct n arbitrary: v w bs cs) apply simp apply (case_tac v rule: bin_exhaust) apply (case_tac w rule: bin_exhaust) apply clarsimp apply (case_tac b) apply auto 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)" apply (induct n arbitrary: v w bs cs) apply simp apply (case_tac v rule: bin_exhaust) apply (case_tac w rule: bin_exhaust) apply clarsimp done 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)" apply (induct n arbitrary: v w bs cs) apply simp apply (case_tac v rule: bin_exhaust) apply (case_tac w rule: bin_exhaust) apply clarsimp done 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/Misc_Typedef.thy b/src/HOL/Word/Misc_Typedef.thy --- a/src/HOL/Word/Misc_Typedef.thy +++ b/src/HOL/Word/Misc_Typedef.thy @@ -1,201 +1,201 @@ (* Author: Jeremy Dawson and Gerwin Klein, NICTA Consequences of type definition theorems, and of extended type definition. *) section \Type Definition Theorems\ theory Misc_Typedef imports Main begin section "More lemmas about normal type definitions" lemma tdD1: "type_definition Rep Abs A \ \x. Rep x \ A" and tdD2: "type_definition Rep Abs A \ \x. Abs (Rep x) = x" and tdD3: "type_definition Rep Abs A \ \y. y \ A \ Rep (Abs y) = y" by (auto simp: type_definition_def) lemma td_nat_int: "type_definition int nat (Collect ((\) 0))" unfolding type_definition_def by auto context type_definition begin declare Rep [iff] Rep_inverse [simp] Rep_inject [simp] lemma Abs_eqD: "Abs x = Abs y \ x \ A \ y \ A \ x = y" by (simp add: Abs_inject) lemma Abs_inverse': "r \ A \ Abs r = a \ Rep a = r" by (safe elim!: Abs_inverse) lemma Rep_comp_inverse: "Rep \ f = g \ Abs \ g = f" using Rep_inverse by auto lemma Rep_eqD [elim!]: "Rep x = Rep y \ x = y" by simp lemma Rep_inverse': "Rep a = r \ Abs r = a" by (safe intro!: Rep_inverse) lemma comp_Abs_inverse: "f \ Abs = g \ g \ Rep = f" using Rep_inverse by auto lemma set_Rep: "A = range Rep" proof (rule set_eqI) show "x \ A \ x \ range Rep" for x by (auto dest: Abs_inverse [of x, symmetric]) qed lemma set_Rep_Abs: "A = range (Rep \ Abs)" proof (rule set_eqI) show "x \ A \ x \ range (Rep \ Abs)" for x by (auto dest: Abs_inverse [of x, symmetric]) qed lemma Abs_inj_on: "inj_on Abs A" unfolding inj_on_def by (auto dest: Abs_inject [THEN iffD1]) lemma image: "Abs ` A = UNIV" - by (auto intro!: image_eqI) + by (fact Abs_image) lemmas td_thm = type_definition_axioms lemma fns1: "Rep \ fa = fr \ Rep \ fa \ Abs = Abs \ fr \ Abs \ fr \ Rep = fa" by (auto dest: Rep_comp_inverse elim: comp_Abs_inverse simp: o_assoc) lemmas fns1a = disjI1 [THEN fns1] lemmas fns1b = disjI2 [THEN fns1] lemma fns4: "Rep \ fa \ Abs = fr \ Rep \ fa = fr \ Rep \ fa \ Abs = Abs \ fr" by auto end interpretation nat_int: type_definition int nat "Collect ((\) 0)" by (rule td_nat_int) declare nat_int.Rep_cases [cases del] nat_int.Abs_cases [cases del] nat_int.Rep_induct [induct del] nat_int.Abs_induct [induct del] subsection "Extended form of type definition predicate" lemma td_conds: "norm \ norm = norm \ fr \ norm = norm \ fr \ norm \ fr \ norm = fr \ norm \ norm \ fr \ norm = norm \ fr" apply safe apply (simp_all add: comp_assoc) apply (simp_all add: o_assoc) done lemma fn_comm_power: "fa \ tr = tr \ fr \ fa ^^ n \ tr = tr \ fr ^^ n" apply (rule ext) apply (induct n) apply (auto dest: fun_cong) done lemmas fn_comm_power' = ext [THEN fn_comm_power, THEN fun_cong, unfolded o_def] locale td_ext = type_definition + fixes norm assumes eq_norm: "\x. Rep (Abs x) = norm x" begin lemma Abs_norm [simp]: "Abs (norm x) = Abs x" using eq_norm [of x] by (auto elim: Rep_inverse') lemma td_th: "g \ Abs = f \ f (Rep x) = g x" by (drule comp_Abs_inverse [symmetric]) simp lemma eq_norm': "Rep \ Abs = norm" by (auto simp: eq_norm) lemma norm_Rep [simp]: "norm (Rep x) = Rep x" by (auto simp: eq_norm' intro: td_th) lemmas td = td_thm lemma set_iff_norm: "w \ A \ w = norm w" by (auto simp: set_Rep_Abs eq_norm' eq_norm [symmetric]) lemma inverse_norm: "Abs n = w \ Rep w = norm n" apply (rule iffI) apply (clarsimp simp add: eq_norm) apply (simp add: eq_norm' [symmetric]) done lemma norm_eq_iff: "norm x = norm y \ Abs x = Abs y" by (simp add: eq_norm' [symmetric]) lemma norm_comps: "Abs \ norm = Abs" "norm \ Rep = Rep" "norm \ norm = norm" by (auto simp: eq_norm' [symmetric] o_def) lemmas norm_norm [simp] = norm_comps lemma fns5: "Rep \ fa \ Abs = fr \ fr \ norm = fr \ norm \ fr = fr" by (fold eq_norm') auto text \ following give conditions for converses to \td_fns1\ \<^item> the condition \norm \ fr \ norm = fr \ norm\ says that \fr\ takes normalised arguments to normalised results \<^item> \norm \ fr \ norm = norm \ fr\ says that \fr\ takes norm-equivalent arguments to norm-equivalent results \<^item> \fr \ norm = fr\ says that \fr\ takes norm-equivalent arguments to the same result \<^item> \norm \ fr = fr\ says that \fr\ takes any argument to a normalised result \ lemma fns2: "Abs \ fr \ Rep = fa \ norm \ fr \ norm = fr \ norm \ Rep \ fa = fr \ Rep" apply (fold eq_norm') apply safe prefer 2 apply (simp add: o_assoc) apply (rule ext) apply (drule_tac x="Rep x" in fun_cong) apply auto done lemma fns3: "Abs \ fr \ Rep = fa \ norm \ fr \ norm = norm \ fr \ fa \ Abs = Abs \ fr" apply (fold eq_norm') apply safe prefer 2 apply (simp add: comp_assoc) apply (rule ext) apply (drule_tac f="a \ b" for a b in fun_cong) apply simp done lemma fns: "fr \ norm = norm \ fr \ fa \ Abs = Abs \ fr \ Rep \ fa = fr \ Rep" apply safe apply (frule fns1b) prefer 2 apply (frule fns1a) apply (rule fns3 [THEN iffD1]) prefer 3 apply (rule fns2 [THEN iffD1]) apply (simp_all add: comp_assoc) apply (simp_all add: o_assoc) done lemma range_norm: "range (Rep \ Abs) = A" by (simp add: set_Rep_Abs) end lemmas td_ext_def' = td_ext_def [unfolded type_definition_def td_ext_axioms_def] end diff --git a/src/HOL/Word/Word.thy b/src/HOL/Word/Word.thy --- a/src/HOL/Word/Word.thy +++ b/src/HOL/Word/Word.thy @@ -1,4611 +1,4609 @@ (* Title: HOL/Word/Word.thy Author: Jeremy Dawson and Gerwin Klein, NICTA *) section \A type of finite bit strings\ theory Word imports "HOL-Library.Type_Length" "HOL-Library.Boolean_Algebra" Bits_Int Bits_Z2 Bit_Comprehension Misc_Typedef Misc_Arithmetic begin text \See \<^file>\Word_Examples.thy\ for examples.\ subsection \Type definition\ typedef (overloaded) 'a word = "{(0::int) ..< 2 ^ LENGTH('a::len0)}" morphisms uint Abs_word by auto lemma uint_nonnegative: "0 \ uint w" using word.uint [of w] by simp lemma uint_bounded: "uint w < 2 ^ LENGTH('a)" for w :: "'a::len0 word" using word.uint [of w] by simp lemma uint_idem: "uint w mod 2 ^ LENGTH('a) = uint w" for w :: "'a::len0 word" using uint_nonnegative uint_bounded by (rule mod_pos_pos_trivial) lemma word_uint_eq_iff: "a = b \ uint a = uint b" by (simp add: uint_inject) lemma word_uint_eqI: "uint a = uint b \ a = b" by (simp add: word_uint_eq_iff) definition word_of_int :: "int \ 'a::len0 word" \ \representation of words using unsigned or signed bins, only difference in these is the type class\ where "word_of_int k = Abs_word (k mod 2 ^ LENGTH('a))" lemma uint_word_of_int: "uint (word_of_int k :: 'a::len0 word) = k mod 2 ^ LENGTH('a)" by (auto simp add: word_of_int_def intro: Abs_word_inverse) lemma word_of_int_uint: "word_of_int (uint w) = w" by (simp add: word_of_int_def uint_idem uint_inverse) lemma split_word_all: "(\x::'a::len0 word. PROP P x) \ (\x. PROP P (word_of_int x))" proof fix x :: "'a word" assume "\x. PROP P (word_of_int x)" then have "PROP P (word_of_int (uint x))" . then show "PROP P x" by (simp add: word_of_int_uint) qed subsection \Type conversions and casting\ definition sint :: "'a::len word \ int" \ \treats the most-significant-bit as a sign bit\ where sint_uint: "sint w = sbintrunc (LENGTH('a) - 1) (uint w)" definition unat :: "'a::len0 word \ nat" where "unat w = nat (uint w)" definition uints :: "nat \ int set" \ \the sets of integers representing the words\ where "uints n = range (bintrunc n)" definition sints :: "nat \ int set" where "sints n = range (sbintrunc (n - 1))" lemma uints_num: "uints n = {i. 0 \ i \ i < 2 ^ n}" by (simp add: uints_def range_bintrunc) lemma sints_num: "sints n = {i. - (2 ^ (n - 1)) \ i \ i < 2 ^ (n - 1)}" by (simp add: sints_def range_sbintrunc) definition unats :: "nat \ nat set" where "unats n = {i. i < 2 ^ n}" definition norm_sint :: "nat \ int \ int" where "norm_sint n w = (w + 2 ^ (n - 1)) mod 2 ^ n - 2 ^ (n - 1)" definition scast :: "'a::len word \ 'b::len word" \ \cast a word to a different length\ where "scast w = word_of_int (sint w)" definition ucast :: "'a::len0 word \ 'b::len0 word" where "ucast w = word_of_int (uint w)" instantiation word :: (len0) size begin definition word_size: "size (w :: 'a word) = LENGTH('a)" instance .. end lemma word_size_gt_0 [iff]: "0 < size w" for w :: "'a::len word" by (simp add: word_size) lemmas lens_gt_0 = word_size_gt_0 len_gt_0 lemma lens_not_0 [iff]: - fixes w :: "'a::len word" - shows "size w \ 0" - and "LENGTH('a) \ 0" + \size w \ 0\ for w :: \'a::len word\ by auto definition source_size :: "('a::len0 word \ 'b) \ nat" \ \whether a cast (or other) function is to a longer or shorter length\ where [code del]: "source_size c = (let arb = undefined; x = c arb in size arb)" definition target_size :: "('a \ 'b::len0 word) \ nat" where [code del]: "target_size c = size (c undefined)" definition is_up :: "('a::len0 word \ 'b::len0 word) \ bool" where "is_up c \ source_size c \ target_size c" definition is_down :: "('a::len0 word \ 'b::len0 word) \ bool" where "is_down c \ target_size c \ source_size c" definition of_bl :: "bool list \ 'a::len0 word" where "of_bl bl = word_of_int (bl_to_bin bl)" definition to_bl :: "'a::len0 word \ bool list" where "to_bl w = bin_to_bl (LENGTH('a)) (uint w)" definition word_reverse :: "'a::len0 word \ 'a word" where "word_reverse w = of_bl (rev (to_bl w))" definition word_int_case :: "(int \ 'b) \ 'a::len0 word \ 'b" where "word_int_case f w = f (uint w)" translations "case x of XCONST of_int y \ b" \ "CONST word_int_case (\y. b) x" "case x of (XCONST of_int :: 'a) y \ b" \ "CONST word_int_case (\y. b) x" subsection \Correspondence relation for theorem transfer\ definition cr_word :: "int \ 'a::len0 word \ bool" where "cr_word = (\x y. word_of_int x = y)" lemma Quotient_word: "Quotient (\x y. bintrunc (LENGTH('a)) x = bintrunc (LENGTH('a)) y) word_of_int uint (cr_word :: _ \ 'a::len0 word \ bool)" unfolding Quotient_alt_def cr_word_def by (simp add: no_bintr_alt1 word_of_int_uint) (simp add: word_of_int_def Abs_word_inject) lemma reflp_word: "reflp (\x y. bintrunc (LENGTH('a::len0)) x = bintrunc (LENGTH('a)) y)" by (simp add: reflp_def) setup_lifting Quotient_word reflp_word text \TODO: The next lemma could be generated automatically.\ lemma uint_transfer [transfer_rule]: "(rel_fun pcr_word (=)) (bintrunc (LENGTH('a))) (uint :: 'a::len0 word \ int)" unfolding rel_fun_def word.pcr_cr_eq cr_word_def by (simp add: no_bintr_alt1 uint_word_of_int) subsection \Basic code generation setup\ definition Word :: "int \ 'a::len0 word" where [code_post]: "Word = word_of_int" lemma [code abstype]: "Word (uint w) = w" by (simp add: Word_def word_of_int_uint) declare uint_word_of_int [code abstract] instantiation word :: (len0) equal begin definition equal_word :: "'a word \ 'a word \ bool" where "equal_word k l \ HOL.equal (uint k) (uint l)" instance by standard (simp add: equal equal_word_def word_uint_eq_iff) end notation fcomp (infixl "\>" 60) notation scomp (infixl "\\" 60) instantiation word :: ("{len0, typerep}") random begin definition "random_word i = Random.range i \\ (\k. Pair ( let j = word_of_int (int_of_integer (integer_of_natural k)) :: 'a word in (j, \_::unit. Code_Evaluation.term_of j)))" instance .. end no_notation fcomp (infixl "\>" 60) no_notation scomp (infixl "\\" 60) subsection \Type-definition locale instantiations\ lemmas uint_0 = uint_nonnegative (* FIXME duplicate *) lemmas uint_lt = uint_bounded (* FIXME duplicate *) lemmas uint_mod_same = uint_idem (* FIXME duplicate *) lemma td_ext_uint: "td_ext (uint :: 'a word \ int) word_of_int (uints (LENGTH('a::len0))) (\w::int. w mod 2 ^ LENGTH('a))" apply (unfold td_ext_def') apply (simp add: uints_num word_of_int_def bintrunc_mod2p) apply (simp add: uint_mod_same uint_0 uint_lt word.uint_inverse word.Abs_word_inverse int_mod_lem) done interpretation word_uint: td_ext "uint::'a::len0 word \ int" word_of_int "uints (LENGTH('a::len0))" "\w. w mod 2 ^ LENGTH('a::len0)" by (fact td_ext_uint) lemmas td_uint = word_uint.td_thm lemmas int_word_uint = word_uint.eq_norm lemma td_ext_ubin: "td_ext (uint :: 'a word \ int) word_of_int (uints (LENGTH('a::len0))) (bintrunc (LENGTH('a)))" by (unfold no_bintr_alt1) (fact td_ext_uint) interpretation word_ubin: td_ext "uint::'a::len0 word \ int" word_of_int "uints (LENGTH('a::len0))" "bintrunc (LENGTH('a::len0))" by (fact td_ext_ubin) subsection \Arithmetic operations\ lift_definition word_succ :: "'a::len0 word \ 'a word" is "\x. x + 1" by (auto simp add: bintrunc_mod2p intro: mod_add_cong) lift_definition word_pred :: "'a::len0 word \ 'a word" is "\x. x - 1" by (auto simp add: bintrunc_mod2p intro: mod_diff_cong) instantiation word :: (len0) "{neg_numeral, modulo, comm_monoid_mult, comm_ring}" begin lift_definition zero_word :: "'a word" is "0" . lift_definition one_word :: "'a word" is "1" . lift_definition plus_word :: "'a word \ 'a word \ 'a word" is "(+)" by (auto simp add: bintrunc_mod2p intro: mod_add_cong) lift_definition minus_word :: "'a word \ 'a word \ 'a word" is "(-)" by (auto simp add: bintrunc_mod2p intro: mod_diff_cong) lift_definition uminus_word :: "'a word \ 'a word" is uminus by (auto simp add: bintrunc_mod2p intro: mod_minus_cong) lift_definition times_word :: "'a word \ 'a word \ 'a word" is "(*)" by (auto simp add: bintrunc_mod2p intro: mod_mult_cong) definition word_div_def: "a div b = word_of_int (uint a div uint b)" definition word_mod_def: "a mod b = word_of_int (uint a mod uint b)" instance by standard (transfer, simp add: algebra_simps)+ end quickcheck_generator word constructors: "zero_class.zero :: ('a::len) word", "numeral :: num \ ('a::len) word", "uminus :: ('a::len) word \ ('a::len) word" text \Legacy theorems:\ lemma word_arith_wis [code]: shows word_add_def: "a + b = word_of_int (uint a + uint b)" and word_sub_wi: "a - b = word_of_int (uint a - uint b)" and word_mult_def: "a * b = word_of_int (uint a * uint b)" and word_minus_def: "- a = word_of_int (- uint a)" and word_succ_alt: "word_succ a = word_of_int (uint a + 1)" and word_pred_alt: "word_pred a = word_of_int (uint a - 1)" and word_0_wi: "0 = word_of_int 0" and word_1_wi: "1 = word_of_int 1" unfolding plus_word_def minus_word_def times_word_def uminus_word_def unfolding word_succ_def word_pred_def zero_word_def one_word_def by simp_all lemma wi_homs: shows wi_hom_add: "word_of_int a + word_of_int b = word_of_int (a + b)" and wi_hom_sub: "word_of_int a - word_of_int b = word_of_int (a - b)" and wi_hom_mult: "word_of_int a * word_of_int b = word_of_int (a * b)" and wi_hom_neg: "- word_of_int a = word_of_int (- a)" and wi_hom_succ: "word_succ (word_of_int a) = word_of_int (a + 1)" and wi_hom_pred: "word_pred (word_of_int a) = word_of_int (a - 1)" by (transfer, simp)+ lemmas wi_hom_syms = wi_homs [symmetric] lemmas word_of_int_homs = wi_homs word_0_wi word_1_wi lemmas word_of_int_hom_syms = word_of_int_homs [symmetric] instance word :: (len) comm_ring_1 proof have *: "0 < LENGTH('a)" by (rule len_gt_0) show "(0::'a word) \ 1" by transfer (use * in \auto simp add: gr0_conv_Suc\) qed lemma word_of_nat: "of_nat n = word_of_int (int n)" by (induct n) (auto simp add : word_of_int_hom_syms) lemma word_of_int: "of_int = word_of_int" apply (rule ext) apply (case_tac x rule: int_diff_cases) apply (simp add: word_of_nat wi_hom_sub) done definition udvd :: "'a::len word \ 'a::len word \ bool" (infixl "udvd" 50) where "a udvd b = (\n\0. uint b = n * uint a)" subsection \Ordering\ instantiation word :: (len0) linorder begin definition word_le_def: "a \ b \ uint a \ uint b" definition word_less_def: "a < b \ uint a < uint b" instance by standard (auto simp: word_less_def word_le_def) end definition word_sle :: "'a::len word \ 'a word \ bool" ("(_/ <=s _)" [50, 51] 50) where "a <=s b \ sint a \ sint b" definition word_sless :: "'a::len word \ 'a word \ bool" ("(_/ x <=s y \ x \ y" subsection \Bit-wise operations\ definition shiftl1 :: "'a::len0 word \ 'a word" where "shiftl1 w = word_of_int (uint w BIT False)" definition shiftr1 :: "'a::len0 word \ 'a word" \ \shift right as unsigned or as signed, ie logical or arithmetic\ where "shiftr1 w = word_of_int (bin_rest (uint w))" instantiation word :: (len0) bit_operations begin lift_definition bitNOT_word :: "'a word \ 'a word" is NOT by (metis bin_trunc_not) lift_definition bitAND_word :: "'a word \ 'a word \ 'a word" is \(AND)\ by (metis bin_trunc_and) lift_definition bitOR_word :: "'a word \ 'a word \ 'a word" is \(OR)\ by (metis bin_trunc_or) lift_definition bitXOR_word :: "'a word \ 'a word \ 'a word" is \(XOR)\ by (metis bin_trunc_xor) definition word_test_bit_def: "test_bit a = bin_nth (uint a)" definition word_set_bit_def: "set_bit a n x = word_of_int (bin_sc n x (uint a))" definition word_lsb_def: "lsb a \ bin_last (uint a)" definition "msb a \ bin_sign (sbintrunc (LENGTH('a) - 1) (uint a)) = - 1" definition shiftl_def: "w << n = (shiftl1 ^^ n) w" definition shiftr_def: "w >> n = (shiftr1 ^^ n) w" instance .. end lemma word_msb_def: "msb a \ bin_sign (sint a) = - 1" by (simp add: msb_word_def sint_uint) lemma [code]: shows word_not_def: "NOT (a::'a::len0 word) = word_of_int (NOT (uint a))" and word_and_def: "(a::'a word) AND b = word_of_int (uint a AND uint b)" and word_or_def: "(a::'a word) OR b = word_of_int (uint a OR uint b)" and word_xor_def: "(a::'a word) XOR b = word_of_int (uint a XOR uint b)" by (simp_all add: bitNOT_word_def bitAND_word_def bitOR_word_def bitXOR_word_def) definition setBit :: "'a::len0 word \ nat \ 'a word" where "setBit w n = set_bit w n True" definition clearBit :: "'a::len0 word \ nat \ 'a word" where "clearBit w n = set_bit w n False" subsection \Shift operations\ definition sshiftr1 :: "'a::len word \ 'a word" where "sshiftr1 w = word_of_int (bin_rest (sint w))" definition bshiftr1 :: "bool \ 'a::len word \ 'a word" where "bshiftr1 b w = of_bl (b # butlast (to_bl w))" definition sshiftr :: "'a::len word \ nat \ 'a word" (infixl ">>>" 55) where "w >>> n = (sshiftr1 ^^ n) w" definition mask :: "nat \ 'a::len word" where "mask n = (1 << n) - 1" definition revcast :: "'a::len0 word \ 'b::len0 word" where "revcast w = of_bl (takefill False (LENGTH('b)) (to_bl w))" definition slice1 :: "nat \ 'a::len0 word \ 'b::len0 word" where "slice1 n w = of_bl (takefill False n (to_bl w))" definition slice :: "nat \ 'a::len0 word \ 'b::len0 word" where "slice n w = slice1 (size w - n) w" subsection \Rotation\ 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" definition word_rotr :: "nat \ 'a::len0 word \ 'a::len0 word" where "word_rotr n w = of_bl (rotater n (to_bl w))" definition word_rotl :: "nat \ 'a::len0 word \ 'a::len0 word" where "word_rotl n w = of_bl (rotate n (to_bl w))" definition word_roti :: "int \ 'a::len0 word \ 'a::len0 word" where "word_roti i w = (if i \ 0 then word_rotr (nat i) w else word_rotl (nat (- i)) w)" subsection \Split and cat operations\ definition word_cat :: "'a::len0 word \ 'b::len0 word \ 'c::len0 word" where "word_cat a b = word_of_int (bin_cat (uint a) (LENGTH('b)) (uint b))" definition word_split :: "'a::len0 word \ 'b::len0 word \ 'c::len0 word" where "word_split a = (case bin_split (LENGTH('c)) (uint a) of (u, v) \ (word_of_int u, word_of_int v))" definition word_rcat :: "'a::len0 word list \ 'b::len0 word" where "word_rcat ws = word_of_int (bin_rcat (LENGTH('a)) (map uint ws))" definition word_rsplit :: "'a::len0 word \ 'b::len word list" where "word_rsplit w = map word_of_int (bin_rsplit (LENGTH('b)) (LENGTH('a), uint w))" definition max_word :: "'a::len word" \ \Largest representable machine integer.\ where "max_word = word_of_int (2 ^ LENGTH('a) - 1)" subsection \Theorems about typedefs\ lemma sint_sbintrunc': "sint (word_of_int bin :: 'a word) = sbintrunc (LENGTH('a::len) - 1) bin" by (auto simp: sint_uint word_ubin.eq_norm sbintrunc_bintrunc_lt) lemma uint_sint: "uint w = bintrunc (LENGTH('a)) (sint w)" for w :: "'a::len word" by (auto simp: sint_uint bintrunc_sbintrunc_le) lemma bintr_uint: "LENGTH('a) \ n \ bintrunc n (uint w) = uint w" for w :: "'a::len0 word" apply (subst word_ubin.norm_Rep [symmetric]) apply (simp only: bintrunc_bintrunc_min word_size) apply (simp add: min.absorb2) done lemma wi_bintr: "LENGTH('a::len0) \ n \ word_of_int (bintrunc n w) = (word_of_int w :: 'a word)" by (auto simp: word_ubin.norm_eq_iff [symmetric] min.absorb1) lemma td_ext_sbin: "td_ext (sint :: 'a word \ int) word_of_int (sints (LENGTH('a::len))) (sbintrunc (LENGTH('a) - 1))" apply (unfold td_ext_def' sint_uint) apply (simp add : word_ubin.eq_norm) apply (cases "LENGTH('a)") apply (auto simp add : sints_def) apply (rule sym [THEN trans]) apply (rule word_ubin.Abs_norm) apply (simp only: bintrunc_sbintrunc) apply (drule sym) apply simp done lemma td_ext_sint: "td_ext (sint :: 'a word \ int) word_of_int (sints (LENGTH('a::len))) (\w. (w + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - 2 ^ (LENGTH('a) - 1))" using td_ext_sbin [where ?'a = 'a] by (simp add: no_sbintr_alt2) text \ We do \sint\ before \sbin\, before \sint\ is the user version and interpretations do not produce thm duplicates. I.e. we get the name \word_sint.Rep_eqD\, but not \word_sbin.Req_eqD\, because the latter is the same thm as the former. \ interpretation word_sint: td_ext "sint ::'a::len word \ int" word_of_int "sints (LENGTH('a::len))" "\w. (w + 2^(LENGTH('a::len) - 1)) mod 2^LENGTH('a::len) - 2 ^ (LENGTH('a::len) - 1)" by (rule td_ext_sint) interpretation word_sbin: td_ext "sint ::'a::len word \ int" word_of_int "sints (LENGTH('a::len))" "sbintrunc (LENGTH('a::len) - 1)" by (rule td_ext_sbin) lemmas int_word_sint = td_ext_sint [THEN td_ext.eq_norm] lemmas td_sint = word_sint.td lemma to_bl_def': "(to_bl :: 'a::len0 word \ bool list) = bin_to_bl (LENGTH('a)) \ uint" by (auto simp: to_bl_def) lemmas word_reverse_no_def [simp] = word_reverse_def [of "numeral w"] for w lemma uints_mod: "uints n = range (\w. w mod 2 ^ n)" by (fact uints_def [unfolded no_bintr_alt1]) lemma word_numeral_alt: "numeral b = word_of_int (numeral b)" by (induct b, simp_all only: numeral.simps word_of_int_homs) declare word_numeral_alt [symmetric, code_abbrev] lemma word_neg_numeral_alt: "- numeral b = word_of_int (- numeral b)" by (simp only: word_numeral_alt wi_hom_neg) declare word_neg_numeral_alt [symmetric, code_abbrev] lemma word_numeral_transfer [transfer_rule]: "(rel_fun (=) pcr_word) numeral numeral" "(rel_fun (=) pcr_word) (- numeral) (- numeral)" apply (simp_all add: rel_fun_def word.pcr_cr_eq cr_word_def) using word_numeral_alt [symmetric] word_neg_numeral_alt [symmetric] by auto lemma uint_bintrunc [simp]: "uint (numeral bin :: 'a word) = bintrunc (LENGTH('a::len0)) (numeral bin)" unfolding word_numeral_alt by (rule word_ubin.eq_norm) lemma uint_bintrunc_neg [simp]: "uint (- numeral bin :: 'a word) = bintrunc (LENGTH('a::len0)) (- numeral bin)" by (simp only: word_neg_numeral_alt word_ubin.eq_norm) lemma sint_sbintrunc [simp]: "sint (numeral bin :: 'a word) = sbintrunc (LENGTH('a::len) - 1) (numeral bin)" by (simp only: word_numeral_alt word_sbin.eq_norm) lemma sint_sbintrunc_neg [simp]: "sint (- numeral bin :: 'a word) = sbintrunc (LENGTH('a::len) - 1) (- numeral bin)" by (simp only: word_neg_numeral_alt word_sbin.eq_norm) lemma unat_bintrunc [simp]: "unat (numeral bin :: 'a::len0 word) = nat (bintrunc (LENGTH('a)) (numeral bin))" by (simp only: unat_def uint_bintrunc) lemma unat_bintrunc_neg [simp]: "unat (- numeral bin :: 'a::len0 word) = nat (bintrunc (LENGTH('a)) (- numeral bin))" by (simp only: unat_def uint_bintrunc_neg) lemma size_0_eq: "size w = 0 \ v = w" for v w :: "'a::len0 word" apply (unfold word_size) apply (rule word_uint.Rep_eqD) apply (rule box_equals) defer apply (rule word_ubin.norm_Rep)+ apply simp done lemma uint_ge_0 [iff]: "0 \ uint x" for x :: "'a::len0 word" using word_uint.Rep [of x] by (simp add: uints_num) lemma uint_lt2p [iff]: "uint x < 2 ^ LENGTH('a)" for x :: "'a::len0 word" using word_uint.Rep [of x] by (simp add: uints_num) lemma sint_ge: "- (2 ^ (LENGTH('a) - 1)) \ sint x" for x :: "'a::len word" using word_sint.Rep [of x] by (simp add: sints_num) lemma sint_lt: "sint x < 2 ^ (LENGTH('a) - 1)" for x :: "'a::len word" using word_sint.Rep [of x] by (simp add: sints_num) lemma sign_uint_Pls [simp]: "bin_sign (uint x) = 0" by (simp add: sign_Pls_ge_0) lemma uint_m2p_neg: "uint x - 2 ^ LENGTH('a) < 0" for x :: "'a::len0 word" by (simp only: diff_less_0_iff_less uint_lt2p) lemma uint_m2p_not_non_neg: "\ 0 \ uint x - 2 ^ LENGTH('a)" for x :: "'a::len0 word" by (simp only: not_le uint_m2p_neg) lemma lt2p_lem: "LENGTH('a) \ n \ uint w < 2 ^ n" for w :: "'a::len0 word" by (metis bintr_uint bintrunc_mod2p int_mod_lem zless2p) lemma uint_le_0_iff [simp]: "uint x \ 0 \ uint x = 0" by (fact uint_ge_0 [THEN leD, THEN antisym_conv1]) lemma uint_nat: "uint w = int (unat w)" by (auto simp: unat_def) lemma uint_numeral: "uint (numeral b :: 'a::len0 word) = numeral b mod 2 ^ LENGTH('a)" by (simp only: word_numeral_alt int_word_uint) lemma uint_neg_numeral: "uint (- numeral b :: 'a::len0 word) = - numeral b mod 2 ^ LENGTH('a)" by (simp only: word_neg_numeral_alt int_word_uint) lemma unat_numeral: "unat (numeral b :: 'a::len0 word) = numeral b mod 2 ^ LENGTH('a)" apply (unfold unat_def) apply (clarsimp simp only: uint_numeral) apply (rule nat_mod_distrib [THEN trans]) apply (rule zero_le_numeral) apply (simp_all add: nat_power_eq) done lemma sint_numeral: "sint (numeral b :: 'a::len word) = (numeral b + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - 2 ^ (LENGTH('a) - 1)" unfolding word_numeral_alt by (rule int_word_sint) lemma word_of_int_0 [simp, code_post]: "word_of_int 0 = 0" unfolding word_0_wi .. lemma word_of_int_1 [simp, code_post]: "word_of_int 1 = 1" unfolding word_1_wi .. lemma word_of_int_neg_1 [simp]: "word_of_int (- 1) = - 1" by (simp add: wi_hom_syms) lemma word_of_int_numeral [simp] : "(word_of_int (numeral bin) :: 'a::len0 word) = numeral bin" by (simp only: word_numeral_alt) lemma word_of_int_neg_numeral [simp]: "(word_of_int (- numeral bin) :: 'a::len0 word) = - numeral bin" by (simp only: word_numeral_alt wi_hom_syms) lemma word_int_case_wi: "word_int_case f (word_of_int i :: 'b word) = f (i mod 2 ^ LENGTH('b::len0))" by (simp add: word_int_case_def word_uint.eq_norm) lemma word_int_split: "P (word_int_case f x) = (\i. x = (word_of_int i :: 'b::len0 word) \ 0 \ i \ i < 2 ^ LENGTH('b) \ P (f i))" - by (auto simp: word_int_case_def word_uint.eq_norm mod_pos_pos_trivial) + by (auto simp: word_int_case_def word_uint.eq_norm) lemma word_int_split_asm: "P (word_int_case f x) = (\n. x = (word_of_int n :: 'b::len0 word) \ 0 \ n \ n < 2 ^ LENGTH('b::len0) \ \ P (f n))" - by (auto simp: word_int_case_def word_uint.eq_norm mod_pos_pos_trivial) + by (auto simp: word_int_case_def word_uint.eq_norm) lemmas uint_range' = word_uint.Rep [unfolded uints_num mem_Collect_eq] lemmas sint_range' = word_sint.Rep [unfolded One_nat_def sints_num mem_Collect_eq] lemma uint_range_size: "0 \ uint w \ uint w < 2 ^ size w" unfolding word_size by (rule uint_range') lemma sint_range_size: "- (2 ^ (size w - Suc 0)) \ sint w \ sint w < 2 ^ (size w - Suc 0)" unfolding word_size by (rule sint_range') lemma sint_above_size: "2 ^ (size w - 1) \ x \ sint w < x" for w :: "'a::len word" unfolding word_size by (rule less_le_trans [OF sint_lt]) lemma sint_below_size: "x \ - (2 ^ (size w - 1)) \ x \ sint w" for w :: "'a::len word" unfolding word_size by (rule order_trans [OF _ sint_ge]) subsection \Testing bits\ lemma test_bit_eq_iff: "test_bit u = test_bit v \ u = v" for u v :: "'a::len0 word" unfolding word_test_bit_def by (simp add: bin_nth_eq_iff) lemma test_bit_size [rule_format] : "w !! n \ n < size w" for w :: "'a::len0 word" apply (unfold word_test_bit_def) apply (subst word_ubin.norm_Rep [symmetric]) apply (simp only: nth_bintr word_size) apply fast done lemma word_eq_iff: "x = y \ (\nn. n < size u \ u !! n = v !! n) \ u = v" for u :: "'a::len0 word" by (simp add: word_size word_eq_iff) lemma word_eqD: "u = v \ u !! x = v !! x" for u v :: "'a::len0 word" by simp lemma test_bit_bin': "w !! n \ n < size w \ bin_nth (uint w) n" by (simp add: word_test_bit_def word_size nth_bintr [symmetric]) lemmas test_bit_bin = test_bit_bin' [unfolded word_size] lemma bin_nth_uint_imp: "bin_nth (uint w) n \ n < LENGTH('a)" for w :: "'a::len0 word" apply (rule nth_bintr [THEN iffD1, THEN conjunct1]) apply (subst word_ubin.norm_Rep) apply assumption done lemma bin_nth_sint: "LENGTH('a) \ n \ bin_nth (sint w) n = bin_nth (sint w) (LENGTH('a) - 1)" for w :: "'a::len word" apply (subst word_sbin.norm_Rep [symmetric]) apply (auto simp add: nth_sbintr) done \ \type definitions theorem for in terms of equivalent bool list\ lemma td_bl: "type_definition (to_bl :: 'a::len0 word \ bool list) of_bl {bl. length bl = LENGTH('a)}" apply (unfold type_definition_def of_bl_def to_bl_def) apply (simp add: word_ubin.eq_norm) apply safe apply (drule sym) apply simp done interpretation word_bl: type_definition "to_bl :: 'a::len0 word \ bool list" of_bl "{bl. length bl = LENGTH('a::len0)}" by (fact td_bl) lemmas word_bl_Rep' = word_bl.Rep [unfolded mem_Collect_eq, iff] lemma word_size_bl: "size w = size (to_bl w)" by (auto simp: word_size) lemma to_bl_use_of_bl: "to_bl w = bl \ w = of_bl bl \ length bl = length (to_bl w)" by (fastforce elim!: word_bl.Abs_inverse [unfolded mem_Collect_eq]) lemma to_bl_word_rev: "to_bl (word_reverse w) = rev (to_bl w)" by (simp add: word_reverse_def word_bl.Abs_inverse) lemma word_rev_rev [simp] : "word_reverse (word_reverse w) = w" by (simp add: word_reverse_def word_bl.Abs_inverse) lemma word_rev_gal: "word_reverse w = u \ word_reverse u = w" by (metis word_rev_rev) lemma word_rev_gal': "u = word_reverse w \ w = word_reverse u" by simp lemma length_bl_gt_0 [iff]: "0 < length (to_bl x)" for x :: "'a::len word" unfolding word_bl_Rep' by (rule len_gt_0) lemma bl_not_Nil [iff]: "to_bl x \ []" for x :: "'a::len word" by (fact length_bl_gt_0 [unfolded length_greater_0_conv]) lemma length_bl_neq_0 [iff]: "length (to_bl x) \ 0" for x :: "'a::len word" by (fact length_bl_gt_0 [THEN gr_implies_not0]) lemma hd_bl_sign_sint: "hd (to_bl w) = (bin_sign (sint w) = -1)" apply (unfold to_bl_def sint_uint) apply (rule trans [OF _ bl_sbin_sign]) apply simp done lemma of_bl_drop': "lend = length bl - LENGTH('a::len0) \ of_bl (drop lend bl) = (of_bl bl :: 'a word)" by (auto simp: of_bl_def trunc_bl2bin [symmetric]) lemma test_bit_of_bl: "(of_bl bl::'a::len0 word) !! n = (rev bl ! n \ n < LENGTH('a) \ n < length bl)" by (auto simp add: of_bl_def word_test_bit_def word_size word_ubin.eq_norm nth_bintr bin_nth_of_bl) lemma no_of_bl: "(numeral bin ::'a::len0 word) = of_bl (bin_to_bl (LENGTH('a)) (numeral bin))" by (simp add: of_bl_def) lemma uint_bl: "to_bl w = bin_to_bl (size w) (uint w)" by (auto simp: word_size to_bl_def) lemma to_bl_bin: "bl_to_bin (to_bl w) = uint w" by (simp add: uint_bl word_size) lemma to_bl_of_bin: "to_bl (word_of_int bin::'a::len0 word) = bin_to_bl (LENGTH('a)) bin" by (auto simp: uint_bl word_ubin.eq_norm word_size) lemma to_bl_numeral [simp]: "to_bl (numeral bin::'a::len0 word) = bin_to_bl (LENGTH('a)) (numeral bin)" unfolding word_numeral_alt by (rule to_bl_of_bin) lemma to_bl_neg_numeral [simp]: "to_bl (- numeral bin::'a::len0 word) = bin_to_bl (LENGTH('a)) (- numeral bin)" unfolding word_neg_numeral_alt by (rule to_bl_of_bin) lemma to_bl_to_bin [simp] : "bl_to_bin (to_bl w) = uint w" by (simp add: uint_bl word_size) lemma uint_bl_bin: "bl_to_bin (bin_to_bl (LENGTH('a)) (uint x)) = uint x" for x :: "'a::len0 word" by (rule trans [OF bin_bl_bin word_ubin.norm_Rep]) \ \naturals\ lemma uints_unats: "uints n = int ` unats n" apply (unfold unats_def uints_num) apply safe apply (rule_tac image_eqI) apply (erule_tac nat_0_le [symmetric]) by auto lemma unats_uints: "unats n = nat ` uints n" by (auto simp: uints_unats image_iff) lemmas bintr_num = word_ubin.norm_eq_iff [of "numeral a" "numeral b", symmetric, folded word_numeral_alt] for a b lemmas sbintr_num = word_sbin.norm_eq_iff [of "numeral a" "numeral b", symmetric, folded word_numeral_alt] for a b lemma num_of_bintr': "bintrunc (LENGTH('a::len0)) (numeral a) = (numeral b) \ numeral a = (numeral b :: 'a word)" unfolding bintr_num by (erule subst, simp) lemma num_of_sbintr': "sbintrunc (LENGTH('a::len) - 1) (numeral a) = (numeral b) \ numeral a = (numeral b :: 'a word)" unfolding sbintr_num by (erule subst, simp) lemma num_abs_bintr: "(numeral x :: 'a word) = word_of_int (bintrunc (LENGTH('a::len0)) (numeral x))" by (simp only: word_ubin.Abs_norm word_numeral_alt) lemma num_abs_sbintr: "(numeral x :: 'a word) = word_of_int (sbintrunc (LENGTH('a::len) - 1) (numeral x))" by (simp only: word_sbin.Abs_norm word_numeral_alt) text \ \cast\ -- note, no arg for new length, as it's determined by type of result, thus in \cast w = w\, the type means cast to length of \w\! \ lemma ucast_id: "ucast w = w" by (auto simp: ucast_def) lemma scast_id: "scast w = w" by (auto simp: scast_def) lemma ucast_bl: "ucast w = of_bl (to_bl w)" by (auto simp: ucast_def of_bl_def uint_bl word_size) lemma nth_ucast: "(ucast w::'a::len0 word) !! n = (w !! n \ n < LENGTH('a))" by (simp add: ucast_def test_bit_bin word_ubin.eq_norm nth_bintr word_size) (fast elim!: bin_nth_uint_imp) \ \literal u(s)cast\ lemma ucast_bintr [simp]: "ucast (numeral w :: 'a::len0 word) = word_of_int (bintrunc (LENGTH('a)) (numeral w))" by (simp add: ucast_def) (* TODO: neg_numeral *) lemma scast_sbintr [simp]: "scast (numeral w ::'a::len word) = word_of_int (sbintrunc (LENGTH('a) - Suc 0) (numeral w))" by (simp add: scast_def) lemma source_size: "source_size (c::'a::len0 word \ _) = LENGTH('a)" unfolding source_size_def word_size Let_def .. lemma target_size: "target_size (c::_ \ 'b::len0 word) = LENGTH('b)" unfolding target_size_def word_size Let_def .. lemma is_down: "is_down c \ LENGTH('b) \ LENGTH('a)" for c :: "'a::len0 word \ 'b::len0 word" by (simp only: is_down_def source_size target_size) lemma is_up: "is_up c \ LENGTH('a) \ LENGTH('b)" for c :: "'a::len0 word \ 'b::len0 word" by (simp only: is_up_def source_size target_size) lemmas is_up_down = trans [OF is_up is_down [symmetric]] lemma down_cast_same [OF refl]: "uc = ucast \ is_down uc \ uc = scast" apply (unfold is_down) apply safe apply (rule ext) apply (unfold ucast_def scast_def uint_sint) apply (rule word_ubin.norm_eq_iff [THEN iffD1]) apply simp done lemma word_rev_tf: "to_bl (of_bl bl::'a::len0 word) = rev (takefill False (LENGTH('a)) (rev bl))" by (auto simp: of_bl_def uint_bl bl_bin_bl_rtf word_ubin.eq_norm word_size) lemma word_rep_drop: "to_bl (of_bl bl::'a::len0 word) = replicate (LENGTH('a) - length bl) False @ drop (length bl - LENGTH('a)) bl" by (simp add: word_rev_tf takefill_alt rev_take) lemma to_bl_ucast: "to_bl (ucast (w::'b::len0 word) ::'a::len0 word) = replicate (LENGTH('a) - LENGTH('b)) False @ drop (LENGTH('b) - LENGTH('a)) (to_bl w)" apply (unfold ucast_bl) apply (rule trans) apply (rule word_rep_drop) apply simp done lemma ucast_up_app [OF refl]: "uc = ucast \ source_size uc + n = target_size uc \ to_bl (uc w) = replicate n False @ (to_bl w)" by (auto simp add : source_size target_size to_bl_ucast) lemma ucast_down_drop [OF refl]: "uc = ucast \ source_size uc = target_size uc + n \ to_bl (uc w) = drop n (to_bl w)" by (auto simp add : source_size target_size to_bl_ucast) lemma scast_down_drop [OF refl]: "sc = scast \ source_size sc = target_size sc + n \ to_bl (sc w) = drop n (to_bl w)" apply (subgoal_tac "sc = ucast") apply safe apply simp apply (erule ucast_down_drop) apply (rule down_cast_same [symmetric]) apply (simp add : source_size target_size is_down) done lemma sint_up_scast [OF refl]: "sc = scast \ is_up sc \ sint (sc w) = sint w" apply (unfold is_up) apply safe apply (simp add: scast_def word_sbin.eq_norm) apply (rule box_equals) prefer 3 apply (rule word_sbin.norm_Rep) apply (rule sbintrunc_sbintrunc_l) defer apply (subst word_sbin.norm_Rep) apply (rule refl) apply simp done lemma uint_up_ucast [OF refl]: "uc = ucast \ is_up uc \ uint (uc w) = uint w" apply (unfold is_up) apply safe apply (rule bin_eqI) apply (fold word_test_bit_def) apply (auto simp add: nth_ucast) apply (auto simp add: test_bit_bin) done lemma ucast_up_ucast [OF refl]: "uc = ucast \ is_up uc \ ucast (uc w) = ucast w" apply (simp (no_asm) add: ucast_def) apply (clarsimp simp add: uint_up_ucast) done lemma scast_up_scast [OF refl]: "sc = scast \ is_up sc \ scast (sc w) = scast w" apply (simp (no_asm) add: scast_def) apply (clarsimp simp add: sint_up_scast) done lemma ucast_of_bl_up [OF refl]: "w = of_bl bl \ size bl \ size w \ ucast w = of_bl bl" by (auto simp add : nth_ucast word_size test_bit_of_bl intro!: word_eqI) lemmas ucast_up_ucast_id = trans [OF ucast_up_ucast ucast_id] lemmas scast_up_scast_id = trans [OF scast_up_scast scast_id] lemmas isduu = is_up_down [where c = "ucast", THEN iffD2] lemmas isdus = is_up_down [where c = "scast", THEN iffD2] lemmas ucast_down_ucast_id = isduu [THEN ucast_up_ucast_id] lemmas scast_down_scast_id = isdus [THEN ucast_up_ucast_id] lemma up_ucast_surj: "is_up (ucast :: 'b::len0 word \ 'a::len0 word) \ surj (ucast :: 'a word \ 'b word)" by (rule surjI) (erule ucast_up_ucast_id) lemma up_scast_surj: "is_up (scast :: 'b::len word \ 'a::len word) \ surj (scast :: 'a word \ 'b word)" by (rule surjI) (erule scast_up_scast_id) lemma down_scast_inj: "is_down (scast :: 'b::len word \ 'a::len word) \ inj_on (ucast :: 'a word \ 'b word) A" by (rule inj_on_inverseI, erule scast_down_scast_id) lemma down_ucast_inj: "is_down (ucast :: 'b::len0 word \ 'a::len0 word) \ inj_on (ucast :: 'a word \ 'b word) A" by (rule inj_on_inverseI) (erule ucast_down_ucast_id) lemma of_bl_append_same: "of_bl (X @ to_bl w) = w" by (rule word_bl.Rep_eqD) (simp add: word_rep_drop) lemma ucast_down_wi [OF refl]: "uc = ucast \ is_down uc \ uc (word_of_int x) = word_of_int x" apply (unfold is_down) apply (clarsimp simp add: ucast_def word_ubin.eq_norm) apply (rule word_ubin.norm_eq_iff [THEN iffD1]) apply (erule bintrunc_bintrunc_ge) done lemma ucast_down_no [OF refl]: "uc = ucast \ is_down uc \ uc (numeral bin) = numeral bin" unfolding word_numeral_alt by clarify (rule ucast_down_wi) lemma ucast_down_bl [OF refl]: "uc = ucast \ is_down uc \ uc (of_bl bl) = of_bl bl" unfolding of_bl_def by clarify (erule ucast_down_wi) lemmas slice_def' = slice_def [unfolded word_size] lemmas test_bit_def' = word_test_bit_def [THEN fun_cong] lemmas word_log_defs = word_and_def word_or_def word_xor_def word_not_def subsection \Word Arithmetic\ lemma word_less_alt: "a < b \ uint a < uint b" by (fact word_less_def) lemma signed_linorder: "class.linorder word_sle word_sless" by standard (auto simp: word_sle_def word_sless_def) interpretation signed: linorder "word_sle" "word_sless" by (rule signed_linorder) lemma udvdI: "0 \ n \ uint b = n * uint a \ a udvd b" by (auto simp: udvd_def) lemmas word_div_no [simp] = word_div_def [of "numeral a" "numeral b"] for a b lemmas word_mod_no [simp] = word_mod_def [of "numeral a" "numeral b"] for a b lemmas word_less_no [simp] = word_less_def [of "numeral a" "numeral b"] for a b lemmas word_le_no [simp] = word_le_def [of "numeral a" "numeral b"] for a b lemmas word_sless_no [simp] = word_sless_def [of "numeral a" "numeral b"] for a b lemmas word_sle_no [simp] = word_sle_def [of "numeral a" "numeral b"] for a b lemma word_m1_wi: "- 1 = word_of_int (- 1)" by (simp add: word_neg_numeral_alt [of Num.One]) lemma word_0_bl [simp]: "of_bl [] = 0" by (simp add: of_bl_def) lemma word_1_bl: "of_bl [True] = 1" by (simp add: of_bl_def bl_to_bin_def) lemma uint_eq_0 [simp]: "uint 0 = 0" unfolding word_0_wi word_ubin.eq_norm by simp lemma of_bl_0 [simp]: "of_bl (replicate n False) = 0" by (simp add: of_bl_def bl_to_bin_rep_False) lemma to_bl_0 [simp]: "to_bl (0::'a::len0 word) = replicate (LENGTH('a)) False" by (simp add: uint_bl word_size bin_to_bl_zero) lemma uint_0_iff: "uint x = 0 \ x = 0" by (simp add: word_uint_eq_iff) lemma unat_0_iff: "unat x = 0 \ x = 0" by (auto simp: unat_def nat_eq_iff uint_0_iff) lemma unat_0 [simp]: "unat 0 = 0" by (auto simp: unat_def) lemma size_0_same': "size w = 0 \ w = v" for v w :: "'a::len0 word" apply (unfold word_size) apply (rule box_equals) defer apply (rule word_uint.Rep_inverse)+ apply (rule word_ubin.norm_eq_iff [THEN iffD1]) apply simp done lemmas size_0_same = size_0_same' [unfolded word_size] lemmas unat_eq_0 = unat_0_iff lemmas unat_eq_zero = unat_0_iff lemma unat_gt_0: "0 < unat x \ x \ 0" by (auto simp: unat_0_iff [symmetric]) lemma ucast_0 [simp]: "ucast 0 = 0" by (simp add: ucast_def) lemma sint_0 [simp]: "sint 0 = 0" by (simp add: sint_uint) lemma scast_0 [simp]: "scast 0 = 0" by (simp add: scast_def) lemma sint_n1 [simp] : "sint (- 1) = - 1" by (simp only: word_m1_wi word_sbin.eq_norm) simp lemma scast_n1 [simp]: "scast (- 1) = - 1" by (simp add: scast_def) lemma uint_1 [simp]: "uint (1::'a::len word) = 1" by (simp only: word_1_wi word_ubin.eq_norm) (simp add: bintrunc_minus_simps(4)) lemma unat_1 [simp]: "unat (1::'a::len word) = 1" by (simp add: unat_def) lemma ucast_1 [simp]: "ucast (1::'a::len word) = 1" by (simp add: ucast_def) \ \now, to get the weaker results analogous to \word_div\/\mod_def\\ subsection \Transferring goals from words to ints\ lemma word_ths: shows word_succ_p1: "word_succ a = a + 1" and word_pred_m1: "word_pred a = a - 1" and word_pred_succ: "word_pred (word_succ a) = a" and word_succ_pred: "word_succ (word_pred a) = a" and word_mult_succ: "word_succ a * b = b + a * b" by (transfer, simp add: algebra_simps)+ lemma uint_cong: "x = y \ uint x = uint y" by simp lemma uint_word_ariths: fixes a b :: "'a::len0 word" shows "uint (a + b) = (uint a + uint b) mod 2 ^ LENGTH('a::len0)" and "uint (a - b) = (uint a - uint b) mod 2 ^ LENGTH('a)" and "uint (a * b) = uint a * uint b mod 2 ^ LENGTH('a)" and "uint (- a) = - uint a mod 2 ^ LENGTH('a)" and "uint (word_succ a) = (uint a + 1) mod 2 ^ LENGTH('a)" and "uint (word_pred a) = (uint a - 1) mod 2 ^ LENGTH('a)" and "uint (0 :: 'a word) = 0 mod 2 ^ LENGTH('a)" and "uint (1 :: 'a word) = 1 mod 2 ^ LENGTH('a)" by (simp_all add: word_arith_wis [THEN trans [OF uint_cong int_word_uint]]) lemma uint_word_arith_bintrs: fixes a b :: "'a::len0 word" shows "uint (a + b) = bintrunc (LENGTH('a)) (uint a + uint b)" and "uint (a - b) = bintrunc (LENGTH('a)) (uint a - uint b)" and "uint (a * b) = bintrunc (LENGTH('a)) (uint a * uint b)" and "uint (- a) = bintrunc (LENGTH('a)) (- uint a)" and "uint (word_succ a) = bintrunc (LENGTH('a)) (uint a + 1)" and "uint (word_pred a) = bintrunc (LENGTH('a)) (uint a - 1)" and "uint (0 :: 'a word) = bintrunc (LENGTH('a)) 0" and "uint (1 :: 'a word) = bintrunc (LENGTH('a)) 1" by (simp_all add: uint_word_ariths bintrunc_mod2p) lemma sint_word_ariths: fixes a b :: "'a::len word" shows "sint (a + b) = sbintrunc (LENGTH('a) - 1) (sint a + sint b)" and "sint (a - b) = sbintrunc (LENGTH('a) - 1) (sint a - sint b)" and "sint (a * b) = sbintrunc (LENGTH('a) - 1) (sint a * sint b)" and "sint (- a) = sbintrunc (LENGTH('a) - 1) (- sint a)" and "sint (word_succ a) = sbintrunc (LENGTH('a) - 1) (sint a + 1)" and "sint (word_pred a) = sbintrunc (LENGTH('a) - 1) (sint a - 1)" and "sint (0 :: 'a word) = sbintrunc (LENGTH('a) - 1) 0" and "sint (1 :: 'a word) = sbintrunc (LENGTH('a) - 1) 1" apply (simp_all only: word_sbin.inverse_norm [symmetric]) apply (simp_all add: wi_hom_syms) apply transfer apply simp apply transfer apply simp done lemmas uint_div_alt = word_div_def [THEN trans [OF uint_cong int_word_uint]] lemmas uint_mod_alt = word_mod_def [THEN trans [OF uint_cong int_word_uint]] lemma word_pred_0_n1: "word_pred 0 = word_of_int (- 1)" unfolding word_pred_m1 by simp lemma succ_pred_no [simp]: "word_succ (numeral w) = numeral w + 1" "word_pred (numeral w) = numeral w - 1" "word_succ (- numeral w) = - numeral w + 1" "word_pred (- numeral w) = - numeral w - 1" by (simp_all add: word_succ_p1 word_pred_m1) lemma word_sp_01 [simp]: "word_succ (- 1) = 0 \ word_succ 0 = 1 \ word_pred 0 = - 1 \ word_pred 1 = 0" by (simp_all add: word_succ_p1 word_pred_m1) \ \alternative approach to lifting arithmetic equalities\ lemma word_of_int_Ex: "\y. x = word_of_int y" by (rule_tac x="uint x" in exI) simp subsection \Order on fixed-length words\ lemma word_zero_le [simp]: "0 \ y" for y :: "'a::len0 word" unfolding word_le_def by auto lemma word_m1_ge [simp] : "word_pred 0 \ y" (* FIXME: delete *) by (simp only: word_le_def word_pred_0_n1 word_uint.eq_norm m1mod2k) auto lemma word_n1_ge [simp]: "y \ -1" for y :: "'a::len0 word" by (simp only: word_le_def word_m1_wi word_uint.eq_norm m1mod2k) auto lemmas word_not_simps [simp] = word_zero_le [THEN leD] word_m1_ge [THEN leD] word_n1_ge [THEN leD] lemma word_gt_0: "0 < y \ 0 \ y" for y :: "'a::len0 word" by (simp add: less_le) lemmas word_gt_0_no [simp] = word_gt_0 [of "numeral y"] for y lemma word_sless_alt: "a sint a < sint b" by (auto simp add: word_sle_def word_sless_def less_le) lemma word_le_nat_alt: "a \ b \ unat a \ unat b" unfolding unat_def word_le_def by (rule nat_le_eq_zle [symmetric]) simp lemma word_less_nat_alt: "a < b \ unat a < unat b" unfolding unat_def word_less_alt by (rule nat_less_eq_zless [symmetric]) simp lemmas unat_mono = word_less_nat_alt [THEN iffD1] instance word :: (len) wellorder proof fix P :: "'a word \ bool" and a assume *: "(\b. (\a. a < b \ P a) \ P b)" have "wf (measure unat)" .. moreover have "{(a, b :: ('a::len) word). a < b} \ measure unat" by (auto simp add: word_less_nat_alt) ultimately have "wf {(a, b :: ('a::len) word). a < b}" by (rule wf_subset) then show "P a" using * by induction blast qed lemma wi_less: "(word_of_int n < (word_of_int m :: 'a::len0 word)) = (n mod 2 ^ LENGTH('a) < m mod 2 ^ LENGTH('a))" unfolding word_less_alt by (simp add: word_uint.eq_norm) lemma wi_le: "(word_of_int n \ (word_of_int m :: 'a::len0 word)) = (n mod 2 ^ LENGTH('a) \ m mod 2 ^ LENGTH('a))" unfolding word_le_def by (simp add: word_uint.eq_norm) lemma udvd_nat_alt: "a udvd b \ (\n\0. unat b = n * unat a)" apply (unfold udvd_def) apply safe apply (simp add: unat_def nat_mult_distrib) apply (simp add: uint_nat) apply (rule exI) apply safe prefer 2 apply (erule notE) apply (rule refl) apply force done lemma udvd_iff_dvd: "x udvd y \ unat x dvd unat y" unfolding dvd_def udvd_nat_alt by force lemma unat_minus_one: assumes "w \ 0" shows "unat (w - 1) = unat w - 1" proof - have "0 \ uint w" by (fact uint_nonnegative) moreover from assms have "0 \ uint w" by (simp add: uint_0_iff) ultimately have "1 \ uint w" by arith from uint_lt2p [of w] have "uint w - 1 < 2 ^ LENGTH('a)" by arith with \1 \ uint w\ have "(uint w - 1) mod 2 ^ LENGTH('a) = uint w - 1" by (auto intro: mod_pos_pos_trivial) with \1 \ uint w\ have "nat ((uint w - 1) mod 2 ^ LENGTH('a)) = nat (uint w) - 1" by auto then show ?thesis by (simp only: unat_def int_word_uint word_arith_wis mod_diff_right_eq) qed lemma measure_unat: "p \ 0 \ unat (p - 1) < unat p" by (simp add: unat_minus_one) (simp add: unat_0_iff [symmetric]) lemmas uint_add_ge0 [simp] = add_nonneg_nonneg [OF uint_ge_0 uint_ge_0] lemmas uint_mult_ge0 [simp] = mult_nonneg_nonneg [OF uint_ge_0 uint_ge_0] lemma uint_sub_lt2p [simp]: "uint x - uint y < 2 ^ LENGTH('a)" for x :: "'a::len0 word" and y :: "'b::len0 word" using uint_ge_0 [of y] uint_lt2p [of x] by arith subsection \Conditions for the addition (etc) of two words to overflow\ lemma uint_add_lem: "(uint x + uint y < 2 ^ LENGTH('a)) = (uint (x + y) = uint x + uint y)" for x y :: "'a::len0 word" by (unfold uint_word_ariths) (auto intro!: trans [OF _ int_mod_lem]) lemma uint_mult_lem: "(uint x * uint y < 2 ^ LENGTH('a)) = (uint (x * y) = uint x * uint y)" for x y :: "'a::len0 word" by (unfold uint_word_ariths) (auto intro!: trans [OF _ int_mod_lem]) lemma uint_sub_lem: "uint x \ uint y \ uint (x - y) = uint x - uint y" by (auto simp: uint_word_ariths intro!: trans [OF _ int_mod_lem]) lemma uint_add_le: "uint (x + y) \ uint x + uint y" unfolding uint_word_ariths by (metis uint_add_ge0 zmod_le_nonneg_dividend) lemma uint_sub_ge: "uint (x - y) \ uint x - uint y" unfolding uint_word_ariths by (metis int_mod_ge uint_sub_lt2p zless2p) lemma mod_add_if_z: "x < z \ y < z \ 0 \ y \ 0 \ x \ 0 \ z \ (x + y) mod z = (if x + y < z then x + y else x + y - z)" for x y z :: int by (auto intro: int_mod_eq) lemma uint_plus_if': "uint (a + b) = (if uint a + uint b < 2 ^ LENGTH('a) then uint a + uint b else uint a + uint b - 2 ^ LENGTH('a))" for a b :: "'a::len0 word" using mod_add_if_z [of "uint a" _ "uint b"] by (simp add: uint_word_ariths) lemma mod_sub_if_z: "x < z \ y < z \ 0 \ y \ 0 \ x \ 0 \ z \ (x - y) mod z = (if y \ x then x - y else x - y + z)" for x y z :: int by (auto intro: int_mod_eq) lemma uint_sub_if': "uint (a - b) = (if uint b \ uint a then uint a - uint b else uint a - uint b + 2 ^ LENGTH('a))" for a b :: "'a::len0 word" using mod_sub_if_z [of "uint a" _ "uint b"] by (simp add: uint_word_ariths) subsection \Definition of \uint_arith\\ lemma word_of_int_inverse: "word_of_int r = a \ 0 \ r \ r < 2 ^ LENGTH('a) \ uint a = r" for a :: "'a::len0 word" apply (erule word_uint.Abs_inverse' [rotated]) apply (simp add: uints_num) done lemma uint_split: "P (uint x) = (\i. word_of_int i = x \ 0 \ i \ i < 2^LENGTH('a) \ P i)" for x :: "'a::len0 word" apply (fold word_int_case_def) - apply (auto dest!: word_of_int_inverse simp: int_word_uint mod_pos_pos_trivial + apply (auto dest!: word_of_int_inverse simp: int_word_uint split: word_int_split) done lemma uint_split_asm: "P (uint x) = (\i. word_of_int i = x \ 0 \ i \ i < 2^LENGTH('a) \ \ P i)" for x :: "'a::len0 word" by (auto dest!: word_of_int_inverse - simp: int_word_uint mod_pos_pos_trivial + simp: int_word_uint split: uint_split) lemmas uint_splits = uint_split uint_split_asm lemmas uint_arith_simps = word_le_def word_less_alt word_uint.Rep_inject [symmetric] uint_sub_if' uint_plus_if' \ \use this to stop, eg. \2 ^ LENGTH(32)\ being simplified\ lemma power_False_cong: "False \ a ^ b = c ^ d" by auto \ \\uint_arith_tac\: reduce to arithmetic on int, try to solve by arith\ ML \ fun uint_arith_simpset ctxt = ctxt addsimps @{thms uint_arith_simps} delsimps @{thms word_uint.Rep_inject} |> fold Splitter.add_split @{thms if_split_asm} |> fold Simplifier.add_cong @{thms power_False_cong} fun uint_arith_tacs ctxt = let fun arith_tac' n t = Arith_Data.arith_tac ctxt n t handle Cooper.COOPER _ => Seq.empty; in [ clarify_tac ctxt 1, full_simp_tac (uint_arith_simpset ctxt) 1, ALLGOALS (full_simp_tac (put_simpset HOL_ss ctxt |> fold Splitter.add_split @{thms uint_splits} |> fold Simplifier.add_cong @{thms power_False_cong})), rewrite_goals_tac ctxt @{thms word_size}, ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN REPEAT (eresolve_tac ctxt [conjE] n) THEN REPEAT (dresolve_tac ctxt @{thms word_of_int_inverse} n THEN assume_tac ctxt n THEN assume_tac ctxt n)), TRYALL arith_tac' ] end fun uint_arith_tac ctxt = SELECT_GOAL (EVERY (uint_arith_tacs ctxt)) \ method_setup uint_arith = \Scan.succeed (SIMPLE_METHOD' o uint_arith_tac)\ "solving word arithmetic via integers and arith" subsection \More on overflows and monotonicity\ lemma no_plus_overflow_uint_size: "x \ x + y \ uint x + uint y < 2 ^ size x" for x y :: "'a::len0 word" unfolding word_size by uint_arith lemmas no_olen_add = no_plus_overflow_uint_size [unfolded word_size] lemma no_ulen_sub: "x \ x - y \ uint y \ uint x" for x y :: "'a::len0 word" by uint_arith lemma no_olen_add': "x \ y + x \ uint y + uint x < 2 ^ LENGTH('a)" for x y :: "'a::len0 word" by (simp add: ac_simps no_olen_add) lemmas olen_add_eqv = trans [OF no_olen_add no_olen_add' [symmetric]] lemmas uint_plus_simple_iff = trans [OF no_olen_add uint_add_lem] lemmas uint_plus_simple = uint_plus_simple_iff [THEN iffD1] lemmas uint_minus_simple_iff = trans [OF no_ulen_sub uint_sub_lem] lemmas uint_minus_simple_alt = uint_sub_lem [folded word_le_def] lemmas word_sub_le_iff = no_ulen_sub [folded word_le_def] lemmas word_sub_le = word_sub_le_iff [THEN iffD2] lemma word_less_sub1: "x \ 0 \ 1 < x \ 0 < x - 1" for x :: "'a::len word" by uint_arith lemma word_le_sub1: "x \ 0 \ 1 \ x \ 0 \ x - 1" for x :: "'a::len word" by uint_arith lemma sub_wrap_lt: "x < x - z \ x < z" for x z :: "'a::len0 word" by uint_arith lemma sub_wrap: "x \ x - z \ z = 0 \ x < z" for x z :: "'a::len0 word" by uint_arith lemma plus_minus_not_NULL_ab: "x \ ab - c \ c \ ab \ c \ 0 \ x + c \ 0" for x ab c :: "'a::len0 word" by uint_arith lemma plus_minus_no_overflow_ab: "x \ ab - c \ c \ ab \ x \ x + c" for x ab c :: "'a::len0 word" by uint_arith lemma le_minus': "a + c \ b \ a \ a + c \ c \ b - a" for a b c :: "'a::len0 word" by uint_arith lemma le_plus': "a \ b \ c \ b - a \ a + c \ b" for a b c :: "'a::len0 word" by uint_arith lemmas le_plus = le_plus' [rotated] lemmas le_minus = leD [THEN thin_rl, THEN le_minus'] (* FIXME *) lemma word_plus_mono_right: "y \ z \ x \ x + z \ x + y \ x + z" for x y z :: "'a::len0 word" by uint_arith lemma word_less_minus_cancel: "y - x < z - x \ x \ z \ y < z" for x y z :: "'a::len0 word" by uint_arith lemma word_less_minus_mono_left: "y < z \ x \ y \ y - x < z - x" for x y z :: "'a::len0 word" by uint_arith lemma word_less_minus_mono: "a < c \ d < b \ a - b < a \ c - d < c \ a - b < c - d" for a b c d :: "'a::len word" by uint_arith lemma word_le_minus_cancel: "y - x \ z - x \ x \ z \ y \ z" for x y z :: "'a::len0 word" by uint_arith lemma word_le_minus_mono_left: "y \ z \ x \ y \ y - x \ z - x" for x y z :: "'a::len0 word" by uint_arith lemma word_le_minus_mono: "a \ c \ d \ b \ a - b \ a \ c - d \ c \ a - b \ c - d" for a b c d :: "'a::len word" by uint_arith lemma plus_le_left_cancel_wrap: "x + y' < x \ x + y < x \ x + y' < x + y \ y' < y" for x y y' :: "'a::len0 word" by uint_arith lemma plus_le_left_cancel_nowrap: "x \ x + y' \ x \ x + y \ x + y' < x + y \ y' < y" for x y y' :: "'a::len0 word" by uint_arith lemma word_plus_mono_right2: "a \ a + b \ c \ b \ a \ a + c" for a b c :: "'a::len0 word" by uint_arith lemma word_less_add_right: "x < y - z \ z \ y \ x + z < y" for x y z :: "'a::len0 word" by uint_arith lemma word_less_sub_right: "x < y + z \ y \ x \ x - y < z" for x y z :: "'a::len0 word" by uint_arith lemma word_le_plus_either: "x \ y \ x \ z \ y \ y + z \ x \ y + z" for x y z :: "'a::len0 word" by uint_arith lemma word_less_nowrapI: "x < z - k \ k \ z \ 0 < k \ x < x + k" for x z k :: "'a::len0 word" by uint_arith lemma inc_le: "i < m \ i + 1 \ m" for i m :: "'a::len word" by uint_arith lemma inc_i: "1 \ i \ i < m \ 1 \ i + 1 \ i + 1 \ m" for i m :: "'a::len word" by uint_arith lemma udvd_incr_lem: "up < uq \ up = ua + n * uint K \ uq = ua + n' * uint K \ up + uint K \ uq" apply clarsimp apply (drule less_le_mult) apply safe done lemma udvd_incr': "p < q \ uint p = ua + n * uint K \ uint q = ua + n' * uint K \ p + K \ q" apply (unfold word_less_alt word_le_def) apply (drule (2) udvd_incr_lem) apply (erule uint_add_le [THEN order_trans]) done lemma udvd_decr': "p < q \ uint p = ua + n * uint K \ uint q = ua + n' * uint K \ p \ q - K" apply (unfold word_less_alt word_le_def) apply (drule (2) udvd_incr_lem) apply (drule le_diff_eq [THEN iffD2]) apply (erule order_trans) apply (rule uint_sub_ge) done lemmas udvd_incr_lem0 = udvd_incr_lem [where ua=0, unfolded add_0_left] lemmas udvd_incr0 = udvd_incr' [where ua=0, unfolded add_0_left] lemmas udvd_decr0 = udvd_decr' [where ua=0, unfolded add_0_left] lemma udvd_minus_le': "xy < k \ z udvd xy \ z udvd k \ xy \ k - z" apply (unfold udvd_def) apply clarify apply (erule (2) udvd_decr0) done lemma udvd_incr2_K: "p < a + s \ a \ a + s \ K udvd s \ K udvd p - a \ a \ p \ 0 < K \ p \ p + K \ p + K \ a + s" supply [[simproc del: linordered_ring_less_cancel_factor]] apply (unfold udvd_def) apply clarify apply (simp add: uint_arith_simps split: if_split_asm) prefer 2 apply (insert uint_range' [of s])[1] apply arith apply (drule add.commute [THEN xtr1]) apply (simp add: diff_less_eq [symmetric]) apply (drule less_le_mult) apply arith apply simp done \ \links with \rbl\ operations\ lemma word_succ_rbl: "to_bl w = bl \ to_bl (word_succ w) = rev (rbl_succ (rev bl))" apply (unfold word_succ_def) apply clarify apply (simp add: to_bl_of_bin) apply (simp add: to_bl_def rbl_succ) done lemma word_pred_rbl: "to_bl w = bl \ to_bl (word_pred w) = rev (rbl_pred (rev bl))" apply (unfold word_pred_def) apply clarify apply (simp add: to_bl_of_bin) apply (simp add: to_bl_def rbl_pred) done lemma word_add_rbl: "to_bl v = vbl \ to_bl w = wbl \ to_bl (v + w) = rev (rbl_add (rev vbl) (rev wbl))" apply (unfold word_add_def) apply clarify apply (simp add: to_bl_of_bin) apply (simp add: to_bl_def rbl_add) done lemma word_mult_rbl: "to_bl v = vbl \ to_bl w = wbl \ to_bl (v * w) = rev (rbl_mult (rev vbl) (rev wbl))" apply (unfold word_mult_def) apply clarify apply (simp add: to_bl_of_bin) apply (simp add: to_bl_def rbl_mult) done lemma rtb_rbl_ariths: "rev (to_bl w) = ys \ rev (to_bl (word_succ w)) = rbl_succ ys" "rev (to_bl w) = ys \ rev (to_bl (word_pred w)) = rbl_pred ys" "rev (to_bl v) = ys \ rev (to_bl w) = xs \ rev (to_bl (v * w)) = rbl_mult ys xs" "rev (to_bl v) = ys \ rev (to_bl w) = xs \ rev (to_bl (v + w)) = rbl_add ys xs" by (auto simp: rev_swap [symmetric] word_succ_rbl word_pred_rbl word_mult_rbl word_add_rbl) subsection \Arithmetic type class instantiations\ lemmas word_le_0_iff [simp] = word_zero_le [THEN leD, THEN antisym_conv1] lemma word_of_int_nat: "0 \ x \ word_of_int x = of_nat (nat x)" by (simp add: word_of_int) text \ note that \iszero_def\ is only for class \comm_semiring_1_cancel\, which requires word length \\ 1\, ie \'a::len word\ \ lemma iszero_word_no [simp]: "iszero (numeral bin :: 'a::len word) = iszero (bintrunc (LENGTH('a)) (numeral bin))" using word_ubin.norm_eq_iff [where 'a='a, of "numeral bin" 0] by (simp add: iszero_def [symmetric]) text \Use \iszero\ to simplify equalities between word numerals.\ lemmas word_eq_numeral_iff_iszero [simp] = eq_numeral_iff_iszero [where 'a="'a::len word"] subsection \Word and nat\ lemma td_ext_unat [OF refl]: "n = LENGTH('a::len) \ td_ext (unat :: 'a word \ nat) of_nat (unats n) (\i. i mod 2 ^ n)" apply (unfold td_ext_def' unat_def word_of_nat unats_uints) apply (auto intro!: imageI simp add : word_of_int_hom_syms) apply (erule word_uint.Abs_inverse [THEN arg_cong]) apply (simp add: int_word_uint nat_mod_distrib nat_power_eq) done lemmas unat_of_nat = td_ext_unat [THEN td_ext.eq_norm] interpretation word_unat: td_ext "unat::'a::len word \ nat" of_nat "unats (LENGTH('a::len))" "\i. i mod 2 ^ LENGTH('a::len)" by (rule td_ext_unat) lemmas td_unat = word_unat.td_thm lemmas unat_lt2p [iff] = word_unat.Rep [unfolded unats_def mem_Collect_eq] lemma unat_le: "y \ unat z \ y \ unats (LENGTH('a))" for z :: "'a::len word" apply (unfold unats_def) apply clarsimp apply (rule xtrans, rule unat_lt2p, assumption) done lemma word_nchotomy: "\w :: 'a::len word. \n. w = of_nat n \ n < 2 ^ LENGTH('a)" apply (rule allI) apply (rule word_unat.Abs_cases) apply (unfold unats_def) apply auto done lemma of_nat_eq: "of_nat n = w \ (\q. n = unat w + q * 2 ^ LENGTH('a))" for w :: "'a::len word" using mod_div_mult_eq [of n "2 ^ LENGTH('a)", symmetric] by (auto simp add: word_unat.inverse_norm) lemma of_nat_eq_size: "of_nat n = w \ (\q. n = unat w + q * 2 ^ size w)" unfolding word_size by (rule of_nat_eq) lemma of_nat_0: "of_nat m = (0::'a::len word) \ (\q. m = q * 2 ^ LENGTH('a))" by (simp add: of_nat_eq) lemma of_nat_2p [simp]: "of_nat (2 ^ LENGTH('a)) = (0::'a::len word)" by (fact mult_1 [symmetric, THEN iffD2 [OF of_nat_0 exI]]) lemma of_nat_gt_0: "of_nat k \ 0 \ 0 < k" by (cases k) auto lemma of_nat_neq_0: "0 < k \ k < 2 ^ LENGTH('a::len) \ of_nat k \ (0 :: 'a word)" by (auto simp add : of_nat_0) lemma Abs_fnat_hom_add: "of_nat a + of_nat b = of_nat (a + b)" by simp lemma Abs_fnat_hom_mult: "of_nat a * of_nat b = (of_nat (a * b) :: 'a::len word)" by (simp add: word_of_nat wi_hom_mult) lemma Abs_fnat_hom_Suc: "word_succ (of_nat a) = of_nat (Suc a)" by (simp add: word_of_nat wi_hom_succ ac_simps) lemma Abs_fnat_hom_0: "(0::'a::len word) = of_nat 0" by simp lemma Abs_fnat_hom_1: "(1::'a::len word) = of_nat (Suc 0)" by simp lemmas Abs_fnat_homs = Abs_fnat_hom_add Abs_fnat_hom_mult Abs_fnat_hom_Suc Abs_fnat_hom_0 Abs_fnat_hom_1 lemma word_arith_nat_add: "a + b = of_nat (unat a + unat b)" by simp lemma word_arith_nat_mult: "a * b = of_nat (unat a * unat b)" by simp lemma word_arith_nat_Suc: "word_succ a = of_nat (Suc (unat a))" by (subst Abs_fnat_hom_Suc [symmetric]) simp lemma word_arith_nat_div: "a div b = of_nat (unat a div unat b)" by (simp add: word_div_def word_of_nat zdiv_int uint_nat) lemma word_arith_nat_mod: "a mod b = of_nat (unat a mod unat b)" by (simp add: word_mod_def word_of_nat zmod_int uint_nat) lemmas word_arith_nat_defs = word_arith_nat_add word_arith_nat_mult word_arith_nat_Suc Abs_fnat_hom_0 Abs_fnat_hom_1 word_arith_nat_div word_arith_nat_mod lemma unat_cong: "x = y \ unat x = unat y" by simp lemmas unat_word_ariths = word_arith_nat_defs [THEN trans [OF unat_cong unat_of_nat]] lemmas word_sub_less_iff = word_sub_le_iff [unfolded linorder_not_less [symmetric] Not_eq_iff] lemma unat_add_lem: "unat x + unat y < 2 ^ LENGTH('a) \ unat (x + y) = unat x + unat y" for x y :: "'a::len word" by (auto simp: unat_word_ariths intro!: trans [OF _ nat_mod_lem]) lemma unat_mult_lem: "unat x * unat y < 2 ^ LENGTH('a) \ unat (x * y) = unat x * unat y" for x y :: "'a::len word" by (auto simp: unat_word_ariths intro!: trans [OF _ nat_mod_lem]) lemmas unat_plus_if' = trans [OF unat_word_ariths(1) mod_nat_add, simplified] lemma le_no_overflow: "x \ b \ a \ a + b \ x \ a + b" for a b x :: "'a::len0 word" apply (erule order_trans) apply (erule olen_add_eqv [THEN iffD1]) done lemmas un_ui_le = trans [OF word_le_nat_alt [symmetric] word_le_def] lemma unat_sub_if_size: "unat (x - y) = (if unat y \ unat x then unat x - unat y else unat x + 2 ^ size x - unat y)" apply (unfold word_size) apply (simp add: un_ui_le) apply (auto simp add: unat_def uint_sub_if') apply (rule nat_diff_distrib) prefer 3 apply (simp add: algebra_simps) apply (rule nat_diff_distrib [THEN trans]) prefer 3 apply (subst nat_add_distrib) prefer 3 apply (simp add: nat_power_eq) apply auto apply uint_arith done lemmas unat_sub_if' = unat_sub_if_size [unfolded word_size] lemma unat_div: "unat (x div y) = unat x div unat y" for x y :: " 'a::len word" apply (simp add : unat_word_ariths) apply (rule unat_lt2p [THEN xtr7, THEN nat_mod_eq']) apply (rule div_le_dividend) done lemma unat_mod: "unat (x mod y) = unat x mod unat y" for x y :: "'a::len word" apply (clarsimp simp add : unat_word_ariths) apply (cases "unat y") prefer 2 apply (rule unat_lt2p [THEN xtr7, THEN nat_mod_eq']) apply (rule mod_le_divisor) apply auto done lemma uint_div: "uint (x div y) = uint x div uint y" for x y :: "'a::len word" by (simp add: uint_nat unat_div zdiv_int) lemma uint_mod: "uint (x mod y) = uint x mod uint y" for x y :: "'a::len word" by (simp add: uint_nat unat_mod zmod_int) text \Definition of \unat_arith\ tactic\ lemma unat_split: "P (unat x) \ (\n. of_nat n = x \ n < 2^LENGTH('a) \ P n)" for x :: "'a::len word" by (auto simp: unat_of_nat) lemma unat_split_asm: "P (unat x) \ (\n. of_nat n = x \ n < 2^LENGTH('a) \ \ P n)" for x :: "'a::len word" by (auto simp: unat_of_nat) lemmas of_nat_inverse = word_unat.Abs_inverse' [rotated, unfolded unats_def, simplified] lemmas unat_splits = unat_split unat_split_asm lemmas unat_arith_simps = word_le_nat_alt word_less_nat_alt word_unat.Rep_inject [symmetric] unat_sub_if' unat_plus_if' unat_div unat_mod \ \\unat_arith_tac\: tactic to reduce word arithmetic to \nat\, try to solve via \arith\\ ML \ fun unat_arith_simpset ctxt = ctxt addsimps @{thms unat_arith_simps} delsimps @{thms word_unat.Rep_inject} |> fold Splitter.add_split @{thms if_split_asm} |> fold Simplifier.add_cong @{thms power_False_cong} fun unat_arith_tacs ctxt = let fun arith_tac' n t = Arith_Data.arith_tac ctxt n t handle Cooper.COOPER _ => Seq.empty; in [ clarify_tac ctxt 1, full_simp_tac (unat_arith_simpset ctxt) 1, ALLGOALS (full_simp_tac (put_simpset HOL_ss ctxt |> fold Splitter.add_split @{thms unat_splits} |> fold Simplifier.add_cong @{thms power_False_cong})), rewrite_goals_tac ctxt @{thms word_size}, ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN REPEAT (eresolve_tac ctxt [conjE] n) THEN REPEAT (dresolve_tac ctxt @{thms of_nat_inverse} n THEN assume_tac ctxt n)), TRYALL arith_tac' ] end fun unat_arith_tac ctxt = SELECT_GOAL (EVERY (unat_arith_tacs ctxt)) \ method_setup unat_arith = \Scan.succeed (SIMPLE_METHOD' o unat_arith_tac)\ "solving word arithmetic via natural numbers and arith" lemma no_plus_overflow_unat_size: "x \ x + y \ unat x + unat y < 2 ^ size x" for x y :: "'a::len word" unfolding word_size by unat_arith lemmas no_olen_add_nat = no_plus_overflow_unat_size [unfolded word_size] lemmas unat_plus_simple = trans [OF no_olen_add_nat unat_add_lem] lemma word_div_mult: "0 < y \ unat x * unat y < 2 ^ LENGTH('a) \ x * y div y = x" for x y :: "'a::len word" apply unat_arith apply clarsimp apply (subst unat_mult_lem [THEN iffD1]) apply auto done lemma div_lt': "i \ k div x \ unat i * unat x < 2 ^ LENGTH('a)" for i k x :: "'a::len word" apply unat_arith apply clarsimp apply (drule mult_le_mono1) apply (erule order_le_less_trans) apply (rule xtr7 [OF unat_lt2p div_mult_le]) done lemmas div_lt'' = order_less_imp_le [THEN div_lt'] lemma div_lt_mult: "i < k div x \ 0 < x \ i * x < k" for i k x :: "'a::len word" apply (frule div_lt'' [THEN unat_mult_lem [THEN iffD1]]) apply (simp add: unat_arith_simps) apply (drule (1) mult_less_mono1) apply (erule order_less_le_trans) apply (rule div_mult_le) done lemma div_le_mult: "i \ k div x \ 0 < x \ i * x \ k" for i k x :: "'a::len word" apply (frule div_lt' [THEN unat_mult_lem [THEN iffD1]]) apply (simp add: unat_arith_simps) apply (drule mult_le_mono1) apply (erule order_trans) apply (rule div_mult_le) done lemma div_lt_uint': "i \ k div x \ uint i * uint x < 2 ^ LENGTH('a)" for i k x :: "'a::len word" apply (unfold uint_nat) apply (drule div_lt') apply (metis of_nat_less_iff of_nat_mult of_nat_numeral of_nat_power) done lemmas div_lt_uint'' = order_less_imp_le [THEN div_lt_uint'] lemma word_le_exists': "x \ y \ \z. y = x + z \ uint x + uint z < 2 ^ LENGTH('a)" for x y z :: "'a::len0 word" apply (rule exI) apply (rule conjI) apply (rule zadd_diff_inverse) apply uint_arith done lemmas plus_minus_not_NULL = order_less_imp_le [THEN plus_minus_not_NULL_ab] lemmas plus_minus_no_overflow = order_less_imp_le [THEN plus_minus_no_overflow_ab] lemmas mcs = word_less_minus_cancel word_less_minus_mono_left word_le_minus_cancel word_le_minus_mono_left lemmas word_l_diffs = mcs [where y = "w + x", unfolded add_diff_cancel] for w x lemmas word_diff_ls = mcs [where z = "w + x", unfolded add_diff_cancel] for w x lemmas word_plus_mcs = word_diff_ls [where y = "v + x", unfolded add_diff_cancel] for v x lemmas le_unat_uoi = unat_le [THEN word_unat.Abs_inverse] lemmas thd = times_div_less_eq_dividend lemmas uno_simps [THEN le_unat_uoi] = mod_le_divisor div_le_dividend dtle lemma word_mod_div_equality: "(n div b) * b + (n mod b) = n" for n b :: "'a::len word" apply (unfold word_less_nat_alt word_arith_nat_defs) apply (cut_tac y="unat b" in gt_or_eq_0) apply (erule disjE) apply (simp only: div_mult_mod_eq uno_simps Word.word_unat.Rep_inverse) apply simp done lemma word_div_mult_le: "a div b * b \ a" for a b :: "'a::len word" apply (unfold word_le_nat_alt word_arith_nat_defs) apply (cut_tac y="unat b" in gt_or_eq_0) apply (erule disjE) apply (simp only: div_mult_le uno_simps Word.word_unat.Rep_inverse) apply simp done lemma word_mod_less_divisor: "0 < n \ m mod n < n" for m n :: "'a::len word" apply (simp only: word_less_nat_alt word_arith_nat_defs) apply (auto simp: uno_simps) done lemma word_of_int_power_hom: "word_of_int a ^ n = (word_of_int (a ^ n) :: 'a::len word)" by (induct n) (simp_all add: wi_hom_mult [symmetric]) lemma word_arith_power_alt: "a ^ n = (word_of_int (uint a ^ n) :: 'a::len word)" by (simp add : word_of_int_power_hom [symmetric]) lemma of_bl_length_less: "length x = k \ k < LENGTH('a) \ (of_bl x :: 'a::len word) < 2 ^ k" apply (unfold of_bl_def word_less_alt word_numeral_alt) apply safe apply (simp (no_asm) add: word_of_int_power_hom word_uint.eq_norm del: word_of_int_numeral) - apply (simp add: mod_pos_pos_trivial) + apply simp apply (subst mod_pos_pos_trivial) apply (rule bl_to_bin_ge0) apply (rule order_less_trans) apply (rule bl_to_bin_lt2p) apply simp apply (rule bl_to_bin_lt2p) done lemma unatSuc: "1 + n \ 0 \ unat (1 + n) = Suc (unat n)" for n :: "'a::len word" by unat_arith subsection \Cardinality, finiteness of set of words\ instance word :: (len0) finite by standard (simp add: type_definition.univ [OF type_definition_word]) lemma card_word: "CARD('a word) = 2 ^ LENGTH('a::len0)" by (simp add: type_definition.card [OF type_definition_word] nat_power_eq) lemma card_word_size: "CARD('a word) = 2 ^ size x" for x :: "'a::len0 word" unfolding word_size by (rule card_word) subsection \Bitwise Operations on Words\ lemma word_eq_rbl_eq: "x = y \ rev (to_bl x) = rev (to_bl y)" by simp lemmas bin_log_bintrs = bin_trunc_not bin_trunc_xor bin_trunc_and bin_trunc_or \ \following definitions require both arithmetic and bit-wise word operations\ \ \to get \word_no_log_defs\ from \word_log_defs\, using \bin_log_bintrs\\ lemmas wils1 = bin_log_bintrs [THEN word_ubin.norm_eq_iff [THEN iffD1], folded word_ubin.eq_norm, THEN eq_reflection] \ \the binary operations only\ (* BH: why is this needed? *) lemmas word_log_binary_defs = word_and_def word_or_def word_xor_def lemma word_wi_log_defs: "NOT (word_of_int a) = word_of_int (NOT a)" "word_of_int a AND word_of_int b = word_of_int (a AND b)" "word_of_int a OR word_of_int b = word_of_int (a OR b)" "word_of_int a XOR word_of_int b = word_of_int (a XOR b)" by (transfer, rule refl)+ lemma word_no_log_defs [simp]: "NOT (numeral a) = word_of_int (NOT (numeral a))" "NOT (- numeral a) = word_of_int (NOT (- numeral a))" "numeral a AND numeral b = word_of_int (numeral a AND numeral b)" "numeral a AND - numeral b = word_of_int (numeral a AND - numeral b)" "- numeral a AND numeral b = word_of_int (- numeral a AND numeral b)" "- numeral a AND - numeral b = word_of_int (- numeral a AND - numeral b)" "numeral a OR numeral b = word_of_int (numeral a OR numeral b)" "numeral a OR - numeral b = word_of_int (numeral a OR - numeral b)" "- numeral a OR numeral b = word_of_int (- numeral a OR numeral b)" "- numeral a OR - numeral b = word_of_int (- numeral a OR - numeral b)" "numeral a XOR numeral b = word_of_int (numeral a XOR numeral b)" "numeral a XOR - numeral b = word_of_int (numeral a XOR - numeral b)" "- numeral a XOR numeral b = word_of_int (- numeral a XOR numeral b)" "- numeral a XOR - numeral b = word_of_int (- numeral a XOR - numeral b)" by (transfer, rule refl)+ text \Special cases for when one of the arguments equals 1.\ lemma word_bitwise_1_simps [simp]: "NOT (1::'a::len0 word) = -2" "1 AND numeral b = word_of_int (1 AND numeral b)" "1 AND - numeral b = word_of_int (1 AND - numeral b)" "numeral a AND 1 = word_of_int (numeral a AND 1)" "- numeral a AND 1 = word_of_int (- numeral a AND 1)" "1 OR numeral b = word_of_int (1 OR numeral b)" "1 OR - numeral b = word_of_int (1 OR - numeral b)" "numeral a OR 1 = word_of_int (numeral a OR 1)" "- numeral a OR 1 = word_of_int (- numeral a OR 1)" "1 XOR numeral b = word_of_int (1 XOR numeral b)" "1 XOR - numeral b = word_of_int (1 XOR - numeral b)" "numeral a XOR 1 = word_of_int (numeral a XOR 1)" "- numeral a XOR 1 = word_of_int (- numeral a XOR 1)" by (transfer, simp)+ text \Special cases for when one of the arguments equals -1.\ lemma word_bitwise_m1_simps [simp]: "NOT (-1::'a::len0 word) = 0" "(-1::'a::len0 word) AND x = x" "x AND (-1::'a::len0 word) = x" "(-1::'a::len0 word) OR x = -1" "x OR (-1::'a::len0 word) = -1" " (-1::'a::len0 word) XOR x = NOT x" "x XOR (-1::'a::len0 word) = NOT x" by (transfer, simp)+ lemma uint_or: "uint (x OR y) = uint x OR uint y" by (transfer, simp add: bin_trunc_ao) lemma uint_and: "uint (x AND y) = uint x AND uint y" by (transfer, simp add: bin_trunc_ao) lemma test_bit_wi [simp]: "(word_of_int x :: 'a::len0 word) !! n \ n < LENGTH('a) \ bin_nth x n" by (simp add: word_test_bit_def word_ubin.eq_norm nth_bintr) lemma word_test_bit_transfer [transfer_rule]: "(rel_fun pcr_word (rel_fun (=) (=))) (\x n. n < LENGTH('a) \ bin_nth x n) (test_bit :: 'a::len0 word \ _)" unfolding rel_fun_def word.pcr_cr_eq cr_word_def by simp lemma word_ops_nth_size: "n < size x \ (x OR y) !! n = (x !! n | y !! n) \ (x AND y) !! n = (x !! n \ y !! n) \ (x XOR y) !! n = (x !! n \ y !! n) \ (NOT x) !! n = (\ x !! n)" for x :: "'a::len0 word" unfolding word_size by transfer (simp add: bin_nth_ops) lemma word_ao_nth: "(x OR y) !! n = (x !! n | y !! n) \ (x AND y) !! n = (x !! n \ y !! n)" for x :: "'a::len0 word" by transfer (auto simp add: bin_nth_ops) lemma test_bit_numeral [simp]: "(numeral w :: 'a::len0 word) !! n \ n < LENGTH('a) \ bin_nth (numeral w) n" by transfer (rule refl) lemma test_bit_neg_numeral [simp]: "(- numeral w :: 'a::len0 word) !! n \ n < LENGTH('a) \ bin_nth (- numeral w) n" by transfer (rule refl) lemma test_bit_1 [simp]: "(1 :: 'a::len word) !! n \ n = 0" by transfer auto lemma nth_0 [simp]: "\ (0 :: 'a::len0 word) !! n" by transfer simp lemma nth_minus1 [simp]: "(-1 :: 'a::len0 word) !! n \ n < LENGTH('a)" by transfer simp \ \get from commutativity, associativity etc of \int_and\ etc to same for \word_and etc\\ lemmas bwsimps = wi_hom_add word_wi_log_defs lemma word_bw_assocs: "(x AND y) AND z = x AND y AND z" "(x OR y) OR z = x OR y OR z" "(x XOR y) XOR z = x XOR y XOR z" for x :: "'a::len0 word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_bw_comms: "x AND y = y AND x" "x OR y = y OR x" "x XOR y = y XOR x" for x :: "'a::len0 word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_bw_lcs: "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 :: "'a::len0 word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_log_esimps [simp]: "x AND 0 = 0" "x AND -1 = x" "x OR 0 = x" "x OR -1 = -1" "x XOR 0 = x" "x XOR -1 = NOT x" "0 AND x = 0" "-1 AND x = x" "0 OR x = x" "-1 OR x = -1" "0 XOR x = x" "-1 XOR x = NOT x" for x :: "'a::len0 word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_not_dist: "NOT (x OR y) = NOT x AND NOT y" "NOT (x AND y) = NOT x OR NOT y" for x :: "'a::len0 word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_bw_same: "x AND x = x" "x OR x = x" "x XOR x = 0" for x :: "'a::len0 word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_ao_absorbs [simp]: "x AND (y OR x) = x" "x OR y AND x = x" "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 :: "'a::len0 word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_not_not [simp]: "NOT (NOT x) = x" for x :: "'a::len0 word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_ao_dist: "(x OR y) AND z = x AND z OR y AND z" for x :: "'a::len0 word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_oa_dist: "x AND y OR z = (x OR z) AND (y OR z)" for x :: "'a::len0 word" by (auto simp: word_eq_iff word_ops_nth_size [unfolded word_size]) lemma word_add_not [simp]: "x + NOT x = -1" for x :: "'a::len0 word" by transfer (simp add: bin_add_not) lemma word_plus_and_or [simp]: "(x AND y) + (x OR y) = x + y" for x :: "'a::len0 word" by transfer (simp add: plus_and_or) lemma leoa: "w = x OR y \ y = w AND y" for x :: "'a::len0 word" by auto lemma leao: "w' = x' AND y' \ x' = x' OR w'" for x' :: "'a::len0 word" by auto lemma word_ao_equiv: "w = w OR w' \ w' = w AND w'" for w w' :: "'a::len0 word" by (auto intro: leoa leao) lemma le_word_or2: "x \ x OR y" for x y :: "'a::len0 word" by (auto simp: word_le_def uint_or intro: le_int_or) lemmas le_word_or1 = xtr3 [OF word_bw_comms (2) le_word_or2] lemmas word_and_le1 = xtr3 [OF word_ao_absorbs (4) [symmetric] le_word_or2] lemmas word_and_le2 = xtr3 [OF word_ao_absorbs (8) [symmetric] le_word_or2] lemma bl_word_not: "to_bl (NOT w) = map Not (to_bl w)" unfolding to_bl_def word_log_defs bl_not_bin by (simp add: word_ubin.eq_norm) lemma bl_word_xor: "to_bl (v XOR w) = map2 (\) (to_bl v) (to_bl w)" unfolding to_bl_def word_log_defs bl_xor_bin by (simp add: word_ubin.eq_norm) lemma bl_word_or: "to_bl (v OR w) = map2 (\) (to_bl v) (to_bl w)" unfolding to_bl_def word_log_defs bl_or_bin by (simp add: word_ubin.eq_norm) lemma bl_word_and: "to_bl (v AND w) = map2 (\) (to_bl v) (to_bl w)" unfolding to_bl_def word_log_defs bl_and_bin by (simp add: word_ubin.eq_norm) lemma word_lsb_alt: "lsb w = test_bit w 0" for w :: "'a::len0 word" by (auto simp: word_test_bit_def word_lsb_def) lemma word_lsb_1_0 [simp]: "lsb (1::'a::len word) \ \ lsb (0::'b::len0 word)" unfolding word_lsb_def uint_eq_0 uint_1 by simp lemma word_lsb_last: "lsb w = last (to_bl w)" for w :: "'a::len word" apply (unfold word_lsb_def uint_bl bin_to_bl_def) apply (rule_tac bin="uint w" in bin_exhaust) apply (cases "size w") apply auto apply (auto simp add: bin_to_bl_aux_alt) done lemma word_lsb_int: "lsb w \ uint w mod 2 = 1" by (auto simp: word_lsb_def bin_last_def) lemma word_msb_sint: "msb w \ sint w < 0" by (simp only: word_msb_def sign_Min_lt_0) lemma msb_word_of_int: "msb (word_of_int x::'a::len word) = bin_nth x (LENGTH('a) - 1)" by (simp add: word_msb_def word_sbin.eq_norm bin_sign_lem) lemma word_msb_numeral [simp]: "msb (numeral w::'a::len word) = bin_nth (numeral w) (LENGTH('a) - 1)" unfolding word_numeral_alt by (rule msb_word_of_int) lemma word_msb_neg_numeral [simp]: "msb (- numeral w::'a::len word) = bin_nth (- numeral w) (LENGTH('a) - 1)" unfolding word_neg_numeral_alt by (rule msb_word_of_int) lemma word_msb_0 [simp]: "\ msb (0::'a::len word)" by (simp add: word_msb_def) lemma word_msb_1 [simp]: "msb (1::'a::len word) \ LENGTH('a) = 1" unfolding word_1_wi msb_word_of_int eq_iff [where 'a=nat] by (simp add: Suc_le_eq) lemma word_msb_nth: "msb w = bin_nth (uint w) (LENGTH('a) - 1)" for w :: "'a::len word" by (simp add: word_msb_def sint_uint bin_sign_lem) lemma word_msb_alt: "msb w = hd (to_bl w)" for w :: "'a::len word" apply (unfold word_msb_nth uint_bl) apply (subst hd_conv_nth) apply (rule length_greater_0_conv [THEN iffD1]) apply simp apply (simp add : nth_bin_to_bl word_size) done lemma word_set_nth [simp]: "set_bit w n (test_bit w n) = w" for w :: "'a::len0 word" by (auto simp: word_test_bit_def word_set_bit_def) lemma bin_nth_uint': "bin_nth (uint w) n \ rev (bin_to_bl (size w) (uint w)) ! n \ n < size w" apply (unfold word_size) apply (safe elim!: bin_nth_uint_imp) apply (frule bin_nth_uint_imp) apply (fast dest!: bin_nth_bl)+ done lemmas bin_nth_uint = bin_nth_uint' [unfolded word_size] lemma test_bit_bl: "w !! n \ rev (to_bl w) ! n \ n < size w" unfolding to_bl_def word_test_bit_def word_size by (rule bin_nth_uint) lemma to_bl_nth: "n < size w \ to_bl w ! n = w !! (size w - Suc n)" apply (unfold test_bit_bl) apply clarsimp apply (rule trans) apply (rule nth_rev_alt) apply (auto simp add: word_size) done lemma test_bit_set: "(set_bit w n x) !! n \ n < size w \ x" for w :: "'a::len0 word" by (auto simp: word_size word_test_bit_def word_set_bit_def word_ubin.eq_norm nth_bintr) lemma test_bit_set_gen: "test_bit (set_bit w n x) m = (if m = n then n < size w \ x else test_bit w m)" for w :: "'a::len0 word" apply (unfold word_size word_test_bit_def word_set_bit_def) apply (clarsimp simp add: word_ubin.eq_norm nth_bintr bin_nth_sc_gen) apply (auto elim!: test_bit_size [unfolded word_size] simp add: word_test_bit_def [symmetric]) done lemma of_bl_rep_False: "of_bl (replicate n False @ bs) = of_bl bs" by (auto simp: of_bl_def bl_to_bin_rep_F) lemma msb_nth: "msb w = w !! (LENGTH('a) - 1)" for w :: "'a::len word" by (simp add: word_msb_nth word_test_bit_def) lemmas msb0 = len_gt_0 [THEN diff_Suc_less, THEN word_ops_nth_size [unfolded word_size]] lemmas msb1 = msb0 [where i = 0] lemmas word_ops_msb = msb1 [unfolded msb_nth [symmetric, unfolded One_nat_def]] lemmas lsb0 = len_gt_0 [THEN word_ops_nth_size [unfolded word_size]] lemmas word_ops_lsb = lsb0 [unfolded word_lsb_alt] lemma word_set_set_same [simp]: "set_bit (set_bit w n x) n y = set_bit w n y" for w :: "'a::len0 word" by (rule word_eqI) (simp add : test_bit_set_gen word_size) lemma word_set_set_diff: fixes w :: "'a::len0 word" assumes "m \ n" shows "set_bit (set_bit w m x) n y = set_bit (set_bit w n y) m x" by (rule word_eqI) (auto simp: test_bit_set_gen word_size assms) lemma nth_sint: fixes w :: "'a::len word" defines "l \ LENGTH('a)" shows "bin_nth (sint w) n = (if n < l - 1 then w !! n else w !! (l - 1))" unfolding sint_uint l_def by (auto simp: nth_sbintr word_test_bit_def [symmetric]) lemma word_lsb_numeral [simp]: "lsb (numeral bin :: 'a::len word) \ bin_last (numeral bin)" unfolding word_lsb_alt test_bit_numeral by simp lemma word_lsb_neg_numeral [simp]: "lsb (- numeral bin :: 'a::len word) \ bin_last (- numeral bin)" by (simp add: word_lsb_alt) lemma set_bit_word_of_int: "set_bit (word_of_int x) n b = word_of_int (bin_sc n b x)" unfolding word_set_bit_def by (rule word_eqI)(simp add: word_size bin_nth_sc_gen word_ubin.eq_norm nth_bintr) lemma word_set_numeral [simp]: "set_bit (numeral bin::'a::len0 word) n b = word_of_int (bin_sc n b (numeral bin))" unfolding word_numeral_alt by (rule set_bit_word_of_int) lemma word_set_neg_numeral [simp]: "set_bit (- numeral bin::'a::len0 word) n b = word_of_int (bin_sc n b (- numeral bin))" unfolding word_neg_numeral_alt by (rule set_bit_word_of_int) lemma word_set_bit_0 [simp]: "set_bit 0 n b = word_of_int (bin_sc n b 0)" unfolding word_0_wi by (rule set_bit_word_of_int) lemma word_set_bit_1 [simp]: "set_bit 1 n b = word_of_int (bin_sc n b 1)" unfolding word_1_wi by (rule set_bit_word_of_int) lemma setBit_no [simp]: "setBit (numeral bin) n = word_of_int (bin_sc n True (numeral bin))" by (simp add: setBit_def) lemma clearBit_no [simp]: "clearBit (numeral bin) n = word_of_int (bin_sc n False (numeral bin))" by (simp add: clearBit_def) lemma to_bl_n1: "to_bl (-1::'a::len0 word) = replicate (LENGTH('a)) True" apply (rule word_bl.Abs_inverse') apply simp apply (rule word_eqI) apply (clarsimp simp add: word_size) apply (auto simp add: word_bl.Abs_inverse test_bit_bl word_size) done lemma word_msb_n1 [simp]: "msb (-1::'a::len word)" unfolding word_msb_alt to_bl_n1 by simp lemma word_set_nth_iff: "set_bit w n b = w \ w !! n = b \ n \ size w" for w :: "'a::len0 word" apply (rule iffI) apply (rule disjCI) apply (drule word_eqD) apply (erule sym [THEN trans]) apply (simp add: test_bit_set) apply (erule disjE) apply clarsimp apply (rule word_eqI) apply (clarsimp simp add : test_bit_set_gen) apply (drule test_bit_size) apply force done lemma test_bit_2p: "(word_of_int (2 ^ n)::'a::len word) !! m \ m = n \ m < LENGTH('a)" by (auto simp: word_test_bit_def word_ubin.eq_norm nth_bintr nth_2p_bin) lemma nth_w2p: "((2::'a::len word) ^ n) !! m \ m = n \ m < LENGTH('a::len)" by (simp add: test_bit_2p [symmetric] word_of_int [symmetric]) lemma uint_2p: "(0::'a::len word) < 2 ^ n \ uint (2 ^ n::'a::len word) = 2 ^ n" apply (unfold word_arith_power_alt) apply (case_tac "LENGTH('a)") apply clarsimp apply (case_tac "nat") apply clarsimp apply (case_tac "n") apply clarsimp apply clarsimp apply (drule word_gt_0 [THEN iffD1]) apply (safe intro!: word_eqI) apply (auto simp add: nth_2p_bin) apply (erule notE) apply (simp (no_asm_use) add: uint_word_of_int word_size) apply (subst mod_pos_pos_trivial) apply simp apply (rule power_strict_increasing) apply simp_all done lemma word_of_int_2p: "(word_of_int (2 ^ n) :: 'a::len word) = 2 ^ n" by (induct n) (simp_all add: wi_hom_syms) lemma bang_is_le: "x !! m \ 2 ^ m \ x" for x :: "'a::len word" apply (rule xtr3) apply (rule_tac [2] y = "x" in le_word_or2) apply (rule word_eqI) apply (auto simp add: word_ao_nth nth_w2p word_size) done lemma word_clr_le: "w \ set_bit w n False" for w :: "'a::len0 word" apply (unfold word_set_bit_def word_le_def word_ubin.eq_norm) apply (rule order_trans) apply (rule bintr_bin_clr_le) apply simp done lemma word_set_ge: "w \ set_bit w n True" for w :: "'a::len word" apply (unfold word_set_bit_def word_le_def word_ubin.eq_norm) apply (rule order_trans [OF _ bintr_bin_set_ge]) apply simp done lemma set_bit_beyond: "size x \ n \ set_bit x n b = x" for x :: "'a :: len0 word" by (auto intro: word_eqI simp add: test_bit_set_gen word_size) lemma rbl_word_or: "rev (to_bl (x OR y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" by (simp add: zip_rev bl_word_or rev_map) lemma rbl_word_and: "rev (to_bl (x AND y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" by (simp add: zip_rev bl_word_and rev_map) lemma rbl_word_xor: "rev (to_bl (x XOR y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" by (simp add: zip_rev bl_word_xor rev_map) lemma rbl_word_not: "rev (to_bl (NOT x)) = map Not (rev (to_bl x))" by (simp add: bl_word_not rev_map) subsection \Bit comprehension\ instantiation word :: (len0) bit_comprehension begin definition word_set_bits_def: "(BITS n. f n) = of_bl (bl_of_nth LENGTH('a) f)" instance .. end lemmas of_nth_def = word_set_bits_def (* FIXME duplicate *) lemma td_ext_nth [OF refl refl refl, unfolded word_size]: "n = size w \ ofn = set_bits \ [w, ofn g] = l \ td_ext test_bit ofn {f. \i. f i \ i < n} (\h i. h i \ i < n)" for w :: "'a::len0 word" apply (unfold word_size td_ext_def') apply safe apply (rule_tac [3] ext) apply (rule_tac [4] ext) apply (unfold word_size of_nth_def test_bit_bl) apply safe defer apply (clarsimp simp: word_bl.Abs_inverse)+ apply (rule word_bl.Rep_inverse') apply (rule sym [THEN trans]) apply (rule bl_of_nth_nth) apply simp apply (rule bl_of_nth_inj) apply (clarsimp simp add : test_bit_bl word_size) done interpretation test_bit: td_ext "(!!) :: 'a::len0 word \ nat \ bool" set_bits "{f. \i. f i \ i < LENGTH('a::len0)}" "(\h i. h i \ i < LENGTH('a::len0))" by (rule td_ext_nth) lemmas td_nth = test_bit.td_thm lemma set_bits_K_False [simp]: "set_bits (\_. False) = (0 :: 'a :: len0 word)" by (rule word_eqI) (simp add: test_bit.eq_norm) subsection \Shifting, Rotating, and Splitting Words\ lemma shiftl1_wi [simp]: "shiftl1 (word_of_int w) = word_of_int (w BIT False)" unfolding shiftl1_def apply (simp add: word_ubin.norm_eq_iff [symmetric] word_ubin.eq_norm) apply (subst refl [THEN bintrunc_BIT_I, symmetric]) apply (subst bintrunc_bintrunc_min) apply simp done lemma shiftl1_numeral [simp]: "shiftl1 (numeral w) = numeral (Num.Bit0 w)" unfolding word_numeral_alt shiftl1_wi by simp lemma shiftl1_neg_numeral [simp]: "shiftl1 (- numeral w) = - numeral (Num.Bit0 w)" unfolding word_neg_numeral_alt shiftl1_wi by simp lemma shiftl1_0 [simp] : "shiftl1 0 = 0" by (simp add: shiftl1_def) lemma shiftl1_def_u: "shiftl1 w = word_of_int (uint w BIT False)" by (simp only: shiftl1_def) (* FIXME: duplicate *) lemma shiftl1_def_s: "shiftl1 w = word_of_int (sint w BIT False)" by (simp add: shiftl1_def Bit_B0 wi_hom_syms) lemma shiftr1_0 [simp]: "shiftr1 0 = 0" by (simp add: shiftr1_def) lemma sshiftr1_0 [simp]: "sshiftr1 0 = 0" by (simp add: sshiftr1_def) lemma sshiftr1_n1 [simp]: "sshiftr1 (- 1) = - 1" by (simp add: sshiftr1_def) lemma shiftl_0 [simp]: "(0::'a::len0 word) << n = 0" by (induct n) (auto simp: shiftl_def) lemma shiftr_0 [simp]: "(0::'a::len0 word) >> n = 0" by (induct n) (auto simp: shiftr_def) lemma sshiftr_0 [simp]: "0 >>> n = 0" by (induct n) (auto simp: sshiftr_def) lemma sshiftr_n1 [simp]: "-1 >>> n = -1" by (induct n) (auto simp: sshiftr_def) lemma nth_shiftl1: "shiftl1 w !! n \ n < size w \ n > 0 \ w !! (n - 1)" apply (unfold shiftl1_def word_test_bit_def) apply (simp add: nth_bintr word_ubin.eq_norm word_size) apply (cases n) apply auto done lemma nth_shiftl': "(w << m) !! n \ n < size w \ n >= m \ w !! (n - m)" for w :: "'a::len0 word" apply (unfold shiftl_def) apply (induct m arbitrary: n) apply (force elim!: test_bit_size) apply (clarsimp simp add : nth_shiftl1 word_size) apply arith done lemmas nth_shiftl = nth_shiftl' [unfolded word_size] lemma nth_shiftr1: "shiftr1 w !! n = w !! Suc n" apply (unfold shiftr1_def word_test_bit_def) apply (simp add: nth_bintr word_ubin.eq_norm) apply safe apply (drule bin_nth.Suc [THEN iffD2, THEN bin_nth_uint_imp]) apply simp done lemma nth_shiftr: "(w >> m) !! n = w !! (n + m)" for w :: "'a::len0 word" apply (unfold shiftr_def) apply (induct "m" arbitrary: n) apply (auto simp add: nth_shiftr1) done text \ see paper page 10, (1), (2), \shiftr1_def\ is of the form of (1), where \f\ (ie \bin_rest\) takes normal arguments to normal results, thus we get (2) from (1) \ lemma uint_shiftr1: "uint (shiftr1 w) = bin_rest (uint w)" apply (unfold shiftr1_def word_ubin.eq_norm bin_rest_trunc_i) apply (subst bintr_uint [symmetric, OF order_refl]) apply (simp only : bintrunc_bintrunc_l) apply simp done lemma nth_sshiftr1: "sshiftr1 w !! n = (if n = size w - 1 then w !! n else w !! Suc n)" apply (unfold sshiftr1_def word_test_bit_def) apply (simp add: nth_bintr word_ubin.eq_norm bin_nth.Suc [symmetric] word_size del: bin_nth.simps) apply (simp add: nth_bintr uint_sint del : bin_nth.simps) apply (auto simp add: bin_nth_sint) done lemma nth_sshiftr [rule_format] : "\n. sshiftr w m !! n = (n < size w \ (if n + m \ size w then w !! (size w - 1) else w !! (n + m)))" apply (unfold sshiftr_def) apply (induct_tac m) apply (simp add: test_bit_bl) apply (clarsimp simp add: nth_sshiftr1 word_size) apply safe apply arith apply arith apply (erule thin_rl) apply (case_tac n) apply safe apply simp apply simp apply (erule thin_rl) apply (case_tac n) apply safe apply simp apply simp apply arith+ done lemma shiftr1_div_2: "uint (shiftr1 w) = uint w div 2" apply (unfold shiftr1_def bin_rest_def) apply (rule word_uint.Abs_inverse) apply (simp add: uints_num pos_imp_zdiv_nonneg_iff) apply (rule xtr7) prefer 2 apply (rule zdiv_le_dividend) apply auto done lemma sshiftr1_div_2: "sint (sshiftr1 w) = sint w div 2" apply (unfold sshiftr1_def bin_rest_def [symmetric]) apply (simp add: word_sbin.eq_norm) apply (rule trans) defer apply (subst word_sbin.norm_Rep [symmetric]) apply (rule refl) apply (subst word_sbin.norm_Rep [symmetric]) apply (unfold One_nat_def) apply (rule sbintrunc_rest) done lemma shiftr_div_2n: "uint (shiftr w n) = uint w div 2 ^ n" apply (unfold shiftr_def) apply (induct n) apply simp apply (simp add: shiftr1_div_2 mult.commute zdiv_zmult2_eq [symmetric]) done lemma sshiftr_div_2n: "sint (sshiftr w n) = sint w div 2 ^ n" apply (unfold sshiftr_def) apply (induct n) apply simp apply (simp add: sshiftr1_div_2 mult.commute zdiv_zmult2_eq [symmetric]) done subsubsection \shift functions in terms of lists of bools\ lemmas bshiftr1_numeral [simp] = bshiftr1_def [where w="numeral w", unfolded to_bl_numeral] for w lemma bshiftr1_bl: "to_bl (bshiftr1 b w) = b # butlast (to_bl w)" unfolding bshiftr1_def by (rule word_bl.Abs_inverse) simp lemma shiftl1_of_bl: "shiftl1 (of_bl bl) = of_bl (bl @ [False])" by (simp add: of_bl_def bl_to_bin_append) lemma shiftl1_bl: "shiftl1 w = of_bl (to_bl w @ [False])" for w :: "'a::len0 word" proof - have "shiftl1 w = shiftl1 (of_bl (to_bl w))" by simp also have "\ = of_bl (to_bl w @ [False])" by (rule shiftl1_of_bl) finally show ?thesis . qed lemma bl_shiftl1: "to_bl (shiftl1 w) = tl (to_bl w) @ [False]" for w :: "'a::len word" by (simp add: shiftl1_bl word_rep_drop drop_Suc drop_Cons') (fast intro!: Suc_leI) \ \Generalized version of \bl_shiftl1\. Maybe this one should replace it?\ lemma bl_shiftl1': "to_bl (shiftl1 w) = tl (to_bl w @ [False])" by (simp add: shiftl1_bl word_rep_drop drop_Suc del: drop_append) lemma shiftr1_bl: "shiftr1 w = of_bl (butlast (to_bl w))" apply (unfold shiftr1_def uint_bl of_bl_def) apply (simp add: butlast_rest_bin word_size) apply (simp add: bin_rest_trunc [symmetric, unfolded One_nat_def]) done lemma bl_shiftr1: "to_bl (shiftr1 w) = False # butlast (to_bl w)" for w :: "'a::len word" by (simp add: shiftr1_bl word_rep_drop len_gt_0 [THEN Suc_leI]) \ \Generalized version of \bl_shiftr1\. Maybe this one should replace it?\ lemma bl_shiftr1': "to_bl (shiftr1 w) = butlast (False # to_bl w)" apply (rule word_bl.Abs_inverse') apply (simp del: butlast.simps) apply (simp add: shiftr1_bl of_bl_def) done lemma shiftl1_rev: "shiftl1 w = word_reverse (shiftr1 (word_reverse w))" apply (unfold word_reverse_def) apply (rule word_bl.Rep_inverse' [symmetric]) apply (simp add: bl_shiftl1' bl_shiftr1' word_bl.Abs_inverse) done lemma shiftl_rev: "shiftl w n = word_reverse (shiftr (word_reverse w) n)" by (induct n) (auto simp add: shiftl_def shiftr_def shiftl1_rev) lemma rev_shiftl: "word_reverse w << n = word_reverse (w >> n)" by (simp add: shiftl_rev) lemma shiftr_rev: "w >> n = word_reverse (word_reverse w << n)" by (simp add: rev_shiftl) lemma rev_shiftr: "word_reverse w >> n = word_reverse (w << n)" by (simp add: shiftr_rev) lemma bl_sshiftr1: "to_bl (sshiftr1 w) = hd (to_bl w) # butlast (to_bl w)" for w :: "'a::len word" apply (unfold sshiftr1_def uint_bl word_size) apply (simp add: butlast_rest_bin word_ubin.eq_norm) apply (simp add: sint_uint) apply (rule nth_equalityI) apply clarsimp apply clarsimp apply (case_tac i) apply (simp_all add: hd_conv_nth length_0_conv [symmetric] nth_bin_to_bl bin_nth.Suc [symmetric] nth_sbintr del: bin_nth.Suc) apply force apply (rule impI) apply (rule_tac f = "bin_nth (uint w)" in arg_cong) apply simp done lemma drop_shiftr: "drop n (to_bl (w >> n)) = take (size w - n) (to_bl w)" for w :: "'a::len word" apply (unfold shiftr_def) apply (induct n) prefer 2 apply (simp add: drop_Suc bl_shiftr1 butlast_drop [symmetric]) apply (rule butlast_take [THEN trans]) apply (auto simp: word_size) done lemma drop_sshiftr: "drop n (to_bl (w >>> n)) = take (size w - n) (to_bl w)" for w :: "'a::len word" apply (unfold sshiftr_def) apply (induct n) prefer 2 apply (simp add: drop_Suc bl_sshiftr1 butlast_drop [symmetric]) apply (rule butlast_take [THEN trans]) apply (auto simp: word_size) done lemma take_shiftr: "n \ size w \ take n (to_bl (w >> n)) = replicate n False" apply (unfold shiftr_def) apply (induct n) prefer 2 apply (simp add: bl_shiftr1' length_0_conv [symmetric] word_size) apply (rule take_butlast [THEN trans]) apply (auto simp: word_size) done lemma take_sshiftr' [rule_format] : "n \ size w \ hd (to_bl (w >>> n)) = hd (to_bl w) \ take n (to_bl (w >>> n)) = replicate n (hd (to_bl w))" for w :: "'a::len word" apply (unfold sshiftr_def) apply (induct n) prefer 2 apply (simp add: bl_sshiftr1) apply (rule impI) apply (rule take_butlast [THEN trans]) apply (auto simp: word_size) done lemmas hd_sshiftr = take_sshiftr' [THEN conjunct1] lemmas take_sshiftr = take_sshiftr' [THEN conjunct2] lemma atd_lem: "take n xs = t \ drop n xs = d \ xs = t @ d" by (auto intro: append_take_drop_id [symmetric]) lemmas bl_shiftr = atd_lem [OF take_shiftr drop_shiftr] lemmas bl_sshiftr = atd_lem [OF take_sshiftr drop_sshiftr] lemma shiftl_of_bl: "of_bl bl << n = of_bl (bl @ replicate n False)" by (induct n) (auto simp: shiftl_def shiftl1_of_bl replicate_app_Cons_same) lemma shiftl_bl: "w << n = of_bl (to_bl w @ replicate n False)" for w :: "'a::len0 word" proof - have "w << n = of_bl (to_bl w) << n" by simp also have "\ = of_bl (to_bl w @ replicate n False)" by (rule shiftl_of_bl) finally show ?thesis . qed lemmas shiftl_numeral [simp] = shiftl_def [where w="numeral w"] for w lemma bl_shiftl: "to_bl (w << n) = drop n (to_bl w) @ replicate (min (size w) n) False" by (simp add: shiftl_bl word_rep_drop word_size) lemma shiftl_zero_size: "size x \ n \ x << n = 0" for x :: "'a::len0 word" apply (unfold word_size) apply (rule word_eqI) apply (clarsimp simp add: shiftl_bl word_size test_bit_of_bl nth_append) done \ \note -- the following results use \'a::len word < number_ring\\ lemma shiftl1_2t: "shiftl1 w = 2 * w" for w :: "'a::len word" by (simp add: shiftl1_def Bit_def wi_hom_mult [symmetric]) lemma shiftl1_p: "shiftl1 w = w + w" for w :: "'a::len word" by (simp add: shiftl1_2t) lemma shiftl_t2n: "shiftl w n = 2 ^ n * w" for w :: "'a::len word" by (induct n) (auto simp: shiftl_def shiftl1_2t) lemma shiftr1_bintr [simp]: "(shiftr1 (numeral w) :: 'a::len0 word) = word_of_int (bin_rest (bintrunc (LENGTH('a)) (numeral w)))" unfolding shiftr1_def word_numeral_alt by (simp add: word_ubin.eq_norm) lemma sshiftr1_sbintr [simp]: "(sshiftr1 (numeral w) :: 'a::len word) = word_of_int (bin_rest (sbintrunc (LENGTH('a) - 1) (numeral w)))" unfolding sshiftr1_def word_numeral_alt by (simp add: word_sbin.eq_norm) lemma shiftr_no [simp]: (* FIXME: neg_numeral *) "(numeral w::'a::len0 word) >> n = word_of_int ((bin_rest ^^ n) (bintrunc (LENGTH('a)) (numeral w)))" by (rule word_eqI) (auto simp: nth_shiftr nth_rest_power_bin nth_bintr word_size) lemma sshiftr_no [simp]: (* FIXME: neg_numeral *) "(numeral w::'a::len word) >>> n = word_of_int ((bin_rest ^^ n) (sbintrunc (LENGTH('a) - 1) (numeral w)))" apply (rule word_eqI) apply (auto simp: nth_sshiftr nth_rest_power_bin nth_sbintr word_size) apply (subgoal_tac "na + n = LENGTH('a) - Suc 0", simp, simp)+ done lemma shiftr1_bl_of: "length bl \ LENGTH('a) \ shiftr1 (of_bl bl::'a::len0 word) = of_bl (butlast bl)" by (clarsimp simp: shiftr1_def of_bl_def butlast_rest_bl2bin word_ubin.eq_norm trunc_bl2bin) lemma shiftr_bl_of: "length bl \ LENGTH('a) \ (of_bl bl::'a::len0 word) >> n = of_bl (take (length bl - n) bl)" apply (unfold shiftr_def) apply (induct n) apply clarsimp apply clarsimp apply (subst shiftr1_bl_of) apply simp apply (simp add: butlast_take) done lemma shiftr_bl: "x >> n \ of_bl (take (LENGTH('a) - n) (to_bl x))" for x :: "'a::len0 word" using shiftr_bl_of [where 'a='a, of "to_bl x"] by simp lemma msb_shift: "msb w \ w >> (LENGTH('a) - 1) \ 0" for w :: "'a::len word" apply (unfold shiftr_bl word_msb_alt) apply (simp add: word_size Suc_le_eq take_Suc) apply (cases "hd (to_bl w)") apply (auto simp: word_1_bl of_bl_rep_False [where n=1 and bs="[]", simplified]) done lemma zip_replicate: "n \ length ys \ zip (replicate n x) ys = map (\y. (x, y)) ys" apply (induct ys arbitrary: n) apply simp_all apply (case_tac n) apply simp_all done lemma align_lem_or [rule_format] : "\x m. length x = n + m \ length y = n + m \ drop m x = replicate n False \ take m y = replicate m False \ map2 (|) x y = take m x @ drop m y" apply (induct y) apply force apply clarsimp apply (case_tac x) apply force apply (case_tac m) apply auto apply (drule_tac t="length xs" for xs in sym) apply (auto simp: zip_replicate o_def) done lemma align_lem_and [rule_format] : "\x m. length x = n + m \ length y = n + m \ drop m x = replicate n False \ take m y = replicate m False \ map2 (\) x y = replicate (n + m) False" apply (induct y) apply force apply clarsimp apply (case_tac x) apply force apply (case_tac m) apply auto apply (drule_tac t="length xs" for xs in sym) apply (auto simp: zip_replicate o_def map_replicate_const) done lemma aligned_bl_add_size [OF refl]: "size x - n = m \ n \ size x \ drop m (to_bl x) = replicate n False \ take m (to_bl y) = replicate m False \ to_bl (x + y) = take m (to_bl x) @ drop m (to_bl y)" apply (subgoal_tac "x AND y = 0") prefer 2 apply (rule word_bl.Rep_eqD) apply (simp add: bl_word_and) apply (rule align_lem_and [THEN trans]) apply (simp_all add: word_size)[5] apply simp apply (subst word_plus_and_or [symmetric]) apply (simp add : bl_word_or) apply (rule align_lem_or) apply (simp_all add: word_size) done subsubsection \Mask\ lemma nth_mask [OF refl, simp]: "m = mask n \ test_bit m i \ i < n \ i < size m" apply (unfold mask_def test_bit_bl) apply (simp only: word_1_bl [symmetric] shiftl_of_bl) apply (clarsimp simp add: word_size) apply (simp only: of_bl_def mask_lem word_of_int_hom_syms add_diff_cancel2) apply (fold of_bl_def) apply (simp add: word_1_bl) apply (rule test_bit_of_bl [THEN trans, unfolded test_bit_bl word_size]) apply auto done lemma mask_bl: "mask n = of_bl (replicate n True)" by (auto simp add : test_bit_of_bl word_size intro: word_eqI) lemma mask_bin: "mask n = word_of_int (bintrunc n (- 1))" by (auto simp add: nth_bintr word_size intro: word_eqI) lemma and_mask_bintr: "w AND mask n = word_of_int (bintrunc n (uint w))" apply (rule word_eqI) apply (simp add: nth_bintr word_size word_ops_nth_size) apply (auto simp add: test_bit_bin) done lemma and_mask_wi: "word_of_int i AND mask n = word_of_int (bintrunc n i)" by (auto simp add: nth_bintr word_size word_ops_nth_size word_eq_iff) lemma and_mask_wi': "word_of_int i AND mask n = (word_of_int (bintrunc (min LENGTH('a) n) i) :: 'a::len word)" by (auto simp add: nth_bintr word_size word_ops_nth_size word_eq_iff) lemma and_mask_no: "numeral i AND mask n = word_of_int (bintrunc n (numeral i))" unfolding word_numeral_alt by (rule and_mask_wi) lemma bl_and_mask': "to_bl (w AND mask n :: 'a::len word) = replicate (LENGTH('a) - n) False @ drop (LENGTH('a) - n) (to_bl w)" apply (rule nth_equalityI) apply simp apply (clarsimp simp add: to_bl_nth word_size) apply (simp add: word_size word_ops_nth_size) apply (auto simp add: word_size test_bit_bl nth_append nth_rev) done lemma and_mask_mod_2p: "w AND mask n = word_of_int (uint w mod 2 ^ n)" by (simp only: and_mask_bintr bintrunc_mod2p) lemma and_mask_lt_2p: "uint (w AND mask n) < 2 ^ n" apply (simp add: and_mask_bintr word_ubin.eq_norm) apply (simp add: bintrunc_mod2p) apply (rule xtr8) prefer 2 apply (rule pos_mod_bound) apply auto done lemma eq_mod_iff: "0 < n \ b = b mod n \ 0 \ b \ b < n" for b n :: int by (simp add: int_mod_lem eq_sym_conv) lemma mask_eq_iff: "w AND mask n = w \ uint w < 2 ^ n" apply (simp add: and_mask_bintr) apply (simp add: word_ubin.inverse_norm) apply (simp add: eq_mod_iff bintrunc_mod2p min_def) apply (fast intro!: lt2p_lem) done lemma and_mask_dvd: "2 ^ n dvd uint w \ w AND mask n = 0" apply (simp add: dvd_eq_mod_eq_0 and_mask_mod_2p) apply (simp add: word_uint.norm_eq_iff [symmetric] word_of_int_homs del: word_of_int_0) apply (subst word_uint.norm_Rep [symmetric]) apply (simp only: bintrunc_bintrunc_min bintrunc_mod2p [symmetric] min_def) apply auto done lemma and_mask_dvd_nat: "2 ^ n dvd unat w \ w AND mask n = 0" apply (unfold unat_def) apply (rule trans [OF _ and_mask_dvd]) apply (unfold dvd_def) apply auto apply (drule uint_ge_0 [THEN nat_int.Abs_inverse' [simplified], symmetric]) apply simp apply (simp add: nat_mult_distrib nat_power_eq) done lemma word_2p_lem: "n < size w \ w < 2 ^ n = (uint w < 2 ^ n)" for w :: "'a::len word" apply (unfold word_size word_less_alt word_numeral_alt) - apply (auto simp add: word_of_int_power_hom word_uint.eq_norm mod_pos_pos_trivial + apply (auto simp add: word_of_int_power_hom word_uint.eq_norm simp del: word_of_int_numeral) done lemma less_mask_eq: "x < 2 ^ n \ x AND mask n = x" for x :: "'a::len word" apply (unfold word_less_alt word_numeral_alt) apply (clarsimp simp add: and_mask_mod_2p word_of_int_power_hom word_uint.eq_norm simp del: word_of_int_numeral) apply (drule xtr8 [rotated]) apply (rule int_mod_le) apply (auto simp add : mod_pos_pos_trivial) done lemmas mask_eq_iff_w2p = trans [OF mask_eq_iff word_2p_lem [symmetric]] lemmas and_mask_less' = iffD2 [OF word_2p_lem and_mask_lt_2p, simplified word_size] lemma and_mask_less_size: "n < size x \ x AND mask n < 2^n" unfolding word_size by (erule and_mask_less') lemma word_mod_2p_is_mask [OF refl]: "c = 2 ^ n \ c > 0 \ x mod c = x AND mask n" for c x :: "'a::len word" by (auto simp: word_mod_def uint_2p and_mask_mod_2p) lemma mask_eqs: "(a AND mask n) + b AND mask n = a + b AND mask n" "a + (b AND mask n) AND mask n = a + b AND mask n" "(a AND mask n) - b AND mask n = a - b AND mask n" "a - (b AND mask n) AND mask n = a - b AND mask n" "a * (b AND mask n) AND mask n = a * b AND mask n" "(b AND mask n) * a AND mask n = b * a AND mask n" "(a AND mask n) + (b AND mask n) AND mask n = a + b AND mask n" "(a AND mask n) - (b AND mask n) AND mask n = a - b AND mask n" "(a AND mask n) * (b AND mask n) AND mask n = a * b AND mask n" "- (a AND mask n) AND mask n = - a AND mask n" "word_succ (a AND mask n) AND mask n = word_succ a AND mask n" "word_pred (a AND mask n) AND mask n = word_pred a AND mask n" using word_of_int_Ex [where x=a] word_of_int_Ex [where x=b] by (auto simp: and_mask_wi' word_of_int_homs word.abs_eq_iff bintrunc_mod2p mod_simps) lemma mask_power_eq: "(x AND mask n) ^ k AND mask n = x ^ k AND mask n" using word_of_int_Ex [where x=x] by (auto simp: and_mask_wi' word_of_int_power_hom word.abs_eq_iff bintrunc_mod2p mod_simps) lemma mask_full [simp]: "mask LENGTH('a) = (- 1 :: 'a::len word)" by (simp add: mask_def word_size shiftl_zero_size) subsubsection \Revcast\ lemmas revcast_def' = revcast_def [simplified] lemmas revcast_def'' = revcast_def' [simplified word_size] lemmas revcast_no_def [simp] = revcast_def' [where w="numeral w", unfolded word_size] for w lemma to_bl_revcast: "to_bl (revcast w :: 'a::len0 word) = takefill False (LENGTH('a)) (to_bl w)" apply (unfold revcast_def' word_size) apply (rule word_bl.Abs_inverse) apply simp done lemma revcast_rev_ucast [OF refl refl refl]: "cs = [rc, uc] \ rc = revcast (word_reverse w) \ uc = ucast w \ rc = word_reverse uc" apply (unfold ucast_def revcast_def' Let_def word_reverse_def) apply (auto simp: to_bl_of_bin takefill_bintrunc) apply (simp add: word_bl.Abs_inverse word_size) done lemma revcast_ucast: "revcast w = word_reverse (ucast (word_reverse w))" using revcast_rev_ucast [of "word_reverse w"] by simp lemma ucast_revcast: "ucast w = word_reverse (revcast (word_reverse w))" by (fact revcast_rev_ucast [THEN word_rev_gal']) lemma ucast_rev_revcast: "ucast (word_reverse w) = word_reverse (revcast w)" by (fact revcast_ucast [THEN word_rev_gal']) text "linking revcast and cast via shift" lemmas wsst_TYs = source_size target_size word_size lemma revcast_down_uu [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (w >> n)" for w :: "'a::len word" apply (simp add: revcast_def') apply (rule word_bl.Rep_inverse') apply (rule trans, rule ucast_down_drop) prefer 2 apply (rule trans, rule drop_shiftr) apply (auto simp: takefill_alt wsst_TYs) done lemma revcast_down_us [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = ucast (w >>> n)" for w :: "'a::len word" apply (simp add: revcast_def') apply (rule word_bl.Rep_inverse') apply (rule trans, rule ucast_down_drop) prefer 2 apply (rule trans, rule drop_sshiftr) apply (auto simp: takefill_alt wsst_TYs) done lemma revcast_down_su [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (w >> n)" for w :: "'a::len word" apply (simp add: revcast_def') apply (rule word_bl.Rep_inverse') apply (rule trans, rule scast_down_drop) prefer 2 apply (rule trans, rule drop_shiftr) apply (auto simp: takefill_alt wsst_TYs) done lemma revcast_down_ss [OF refl]: "rc = revcast \ source_size rc = target_size rc + n \ rc w = scast (w >>> n)" for w :: "'a::len word" apply (simp add: revcast_def') apply (rule word_bl.Rep_inverse') apply (rule trans, rule scast_down_drop) prefer 2 apply (rule trans, rule drop_sshiftr) apply (auto simp: takefill_alt wsst_TYs) done (* FIXME: should this also be [OF refl] ? *) lemma cast_down_rev: "uc = ucast \ source_size uc = target_size uc + n \ uc w = revcast (w << n)" for w :: "'a::len word" apply (unfold shiftl_rev) apply clarify apply (simp add: revcast_rev_ucast) apply (rule word_rev_gal') apply (rule trans [OF _ revcast_rev_ucast]) apply (rule revcast_down_uu [symmetric]) apply (auto simp add: wsst_TYs) done lemma revcast_up [OF refl]: "rc = revcast \ source_size rc + n = target_size rc \ rc w = (ucast w :: 'a::len word) << n" apply (simp add: revcast_def') apply (rule word_bl.Rep_inverse') apply (simp add: takefill_alt) apply (rule bl_shiftl [THEN trans]) apply (subst ucast_up_app) apply (auto simp add: wsst_TYs) done lemmas rc1 = revcast_up [THEN revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] lemmas rc2 = revcast_down_uu [THEN revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] lemmas ucast_up = rc1 [simplified rev_shiftr [symmetric] revcast_ucast [symmetric]] lemmas ucast_down = rc2 [simplified rev_shiftr revcast_ucast [symmetric]] subsubsection \Slices\ lemma slice1_no_bin [simp]: "slice1 n (numeral w :: 'b word) = of_bl (takefill False n (bin_to_bl (LENGTH('b::len0)) (numeral w)))" by (simp add: slice1_def) (* TODO: neg_numeral *) lemma slice_no_bin [simp]: "slice n (numeral w :: 'b word) = of_bl (takefill False (LENGTH('b::len0) - n) (bin_to_bl (LENGTH('b::len0)) (numeral w)))" by (simp add: slice_def word_size) (* TODO: neg_numeral *) lemma slice1_0 [simp] : "slice1 n 0 = 0" unfolding slice1_def by simp lemma slice_0 [simp] : "slice n 0 = 0" unfolding slice_def by auto lemma slice_take': "slice n w = of_bl (take (size w - n) (to_bl w))" unfolding slice_def' slice1_def by (simp add : takefill_alt word_size) lemmas slice_take = slice_take' [unfolded word_size] \ \shiftr to a word of the same size is just slice, slice is just shiftr then ucast\ lemmas shiftr_slice = trans [OF shiftr_bl [THEN meta_eq_to_obj_eq] slice_take [symmetric]] lemma slice_shiftr: "slice n w = ucast (w >> n)" apply (unfold slice_take shiftr_bl) apply (rule ucast_of_bl_up [symmetric]) apply (simp add: word_size) done lemma nth_slice: "(slice n w :: 'a::len0 word) !! m = (w !! (m + n) \ m < LENGTH('a))" by (simp add: slice_shiftr nth_ucast nth_shiftr) lemma slice1_down_alt': "sl = slice1 n w \ fs = size sl \ fs + k = n \ to_bl sl = takefill False fs (drop k (to_bl w))" by (auto simp: slice1_def word_size of_bl_def uint_bl word_ubin.eq_norm bl_bin_bl_rep_drop drop_takefill) lemma slice1_up_alt': "sl = slice1 n w \ fs = size sl \ fs = n + k \ to_bl sl = takefill False fs (replicate k False @ (to_bl w))" apply (unfold slice1_def word_size of_bl_def uint_bl) apply (clarsimp simp: word_ubin.eq_norm bl_bin_bl_rep_drop takefill_append [symmetric]) apply (rule_tac f = "\k. takefill False (LENGTH('a)) (replicate k False @ bin_to_bl (LENGTH('b)) (uint w))" in arg_cong) apply arith done lemmas sd1 = slice1_down_alt' [OF refl refl, unfolded word_size] lemmas su1 = slice1_up_alt' [OF refl refl, unfolded word_size] lemmas slice1_down_alt = le_add_diff_inverse [THEN sd1] lemmas slice1_up_alts = le_add_diff_inverse [symmetric, THEN su1] le_add_diff_inverse2 [symmetric, THEN su1] lemma ucast_slice1: "ucast w = slice1 (size w) w" by (simp add: slice1_def ucast_bl takefill_same' word_size) lemma ucast_slice: "ucast w = slice 0 w" by (simp add: slice_def ucast_slice1) lemma slice_id: "slice 0 t = t" by (simp only: ucast_slice [symmetric] ucast_id) lemma revcast_slice1 [OF refl]: "rc = revcast w \ slice1 (size rc) w = rc" by (simp add: slice1_def revcast_def' word_size) lemma slice1_tf_tf': "to_bl (slice1 n w :: 'a::len0 word) = rev (takefill False (LENGTH('a)) (rev (takefill False n (to_bl w))))" unfolding slice1_def by (rule word_rev_tf) lemmas slice1_tf_tf = slice1_tf_tf' [THEN word_bl.Rep_inverse', symmetric] lemma rev_slice1: "n + k = LENGTH('a) + LENGTH('b) \ slice1 n (word_reverse w :: 'b::len0 word) = word_reverse (slice1 k w :: 'a::len0 word)" apply (unfold word_reverse_def slice1_tf_tf) apply (rule word_bl.Rep_inverse') apply (rule rev_swap [THEN iffD1]) apply (rule trans [symmetric]) apply (rule tf_rev) apply (simp add: word_bl.Abs_inverse) apply (simp add: word_bl.Abs_inverse) done lemma rev_slice: "n + k + LENGTH('a::len0) = LENGTH('b::len0) \ slice n (word_reverse (w::'b word)) = word_reverse (slice k w :: 'a word)" apply (unfold slice_def word_size) apply (rule rev_slice1) apply arith done lemmas sym_notr = not_iff [THEN iffD2, THEN not_sym, THEN not_iff [THEN iffD1]] \ \problem posed by TPHOLs referee: criterion for overflow of addition of signed integers\ lemma sofl_test: "(sint x + sint y = sint (x + y)) = ((((x + y) XOR x) AND ((x + y) XOR y)) >> (size x - 1) = 0)" for x y :: "'a::len word" apply (unfold word_size) apply (cases "LENGTH('a)", simp) apply (subst msb_shift [THEN sym_notr]) apply (simp add: word_ops_msb) apply (simp add: word_msb_sint) apply safe apply simp_all apply (unfold sint_word_ariths) apply (unfold word_sbin.set_iff_norm [symmetric] sints_num) apply safe apply (insert sint_range' [where x=x]) apply (insert sint_range' [where x=y]) defer apply (simp (no_asm), arith) apply (simp (no_asm), arith) defer defer apply (simp (no_asm), arith) apply (simp (no_asm), arith) apply (rule notI [THEN notnotD], drule leI not_le_imp_less, drule sbintrunc_inc sbintrunc_dec, simp)+ done lemma shiftr_zero_size: "size x \ n \ x >> n = 0" for x :: "'a :: len0 word" by (rule word_eqI) (auto simp add: nth_shiftr dest: test_bit_size) subsection \Split and cat\ lemmas word_split_bin' = word_split_def lemmas word_cat_bin' = word_cat_def lemma word_rsplit_no: "(word_rsplit (numeral bin :: 'b::len0 word) :: 'a word list) = map word_of_int (bin_rsplit (LENGTH('a::len)) (LENGTH('b), bintrunc (LENGTH('b)) (numeral bin)))" by (simp add: word_rsplit_def word_ubin.eq_norm) lemmas word_rsplit_no_cl [simp] = word_rsplit_no [unfolded bin_rsplitl_def bin_rsplit_l [symmetric]] lemma test_bit_cat: "wc = word_cat a b \ wc !! n = (n < size wc \ (if n < size b then b !! n else a !! (n - size b)))" apply (auto simp: word_cat_bin' test_bit_bin word_ubin.eq_norm nth_bintr bin_nth_cat word_size) apply (erule bin_nth_uint_imp) done lemma word_cat_bl: "word_cat a b = of_bl (to_bl a @ to_bl b)" by (simp add: of_bl_def to_bl_def word_cat_bin' bl_to_bin_app_cat) lemma of_bl_append: "(of_bl (xs @ ys) :: 'a::len word) = of_bl xs * 2^(length ys) + of_bl ys" apply (simp add: of_bl_def bl_to_bin_app_cat bin_cat_num) apply (simp add: word_of_int_power_hom [symmetric] word_of_int_hom_syms) done lemma of_bl_False [simp]: "of_bl (False#xs) = of_bl xs" by (rule word_eqI) (auto simp: test_bit_of_bl nth_append) lemma of_bl_True [simp]: "(of_bl (True # xs) :: 'a::len word) = 2^length xs + of_bl xs" by (subst of_bl_append [where xs="[True]", simplified]) (simp add: word_1_bl) lemma of_bl_Cons: "of_bl (x#xs) = of_bool x * 2^length xs + of_bl xs" by (cases x) simp_all lemma split_uint_lem: "bin_split n (uint w) = (a, b) \ a = bintrunc (LENGTH('a) - n) a \ b = bintrunc (LENGTH('a)) b" for w :: "'a::len0 word" apply (frule word_ubin.norm_Rep [THEN ssubst]) apply (drule bin_split_trunc1) apply (drule sym [THEN trans]) apply assumption apply safe done lemma word_split_bl': "std = size c - size b \ (word_split c = (a, b)) \ (a = of_bl (take std (to_bl c)) \ b = of_bl (drop std (to_bl c)))" apply (unfold word_split_bin') apply safe defer apply (clarsimp split: prod.splits) apply hypsubst_thin apply (drule word_ubin.norm_Rep [THEN ssubst]) apply (drule split_bintrunc) apply (simp add: of_bl_def bl2bin_drop word_size word_ubin.norm_eq_iff [symmetric] min_def del: word_ubin.norm_Rep) apply (clarsimp split: prod.splits) apply (frule split_uint_lem [THEN conjunct1]) apply (unfold word_size) apply (cases "LENGTH('a) \ LENGTH('b)") defer apply simp apply (simp add : of_bl_def to_bl_def) apply (subst bin_split_take1 [symmetric]) prefer 2 apply assumption apply simp apply (erule thin_rl) apply (erule arg_cong [THEN trans]) apply (simp add : word_ubin.norm_eq_iff [symmetric]) done lemma word_split_bl: "std = size c - size b \ (a = of_bl (take std (to_bl c)) \ b = of_bl (drop std (to_bl c))) \ word_split c = (a, b)" apply (rule iffI) defer apply (erule (1) word_split_bl') apply (case_tac "word_split c") apply (auto simp add: word_size) apply (frule word_split_bl' [rotated]) apply (auto simp add: word_size) done lemma word_split_bl_eq: "(word_split c :: ('c::len0 word \ 'd::len0 word)) = (of_bl (take (LENGTH('a::len) - LENGTH('d::len0)) (to_bl c)), of_bl (drop (LENGTH('a) - LENGTH('d)) (to_bl c)))" for c :: "'a::len word" apply (rule word_split_bl [THEN iffD1]) apply (unfold word_size) apply (rule refl conjI)+ done \ \keep quantifiers for use in simplification\ lemma test_bit_split': "word_split c = (a, b) \ (\n m. b !! n = (n < size b \ c !! n) \ a !! m = (m < size a \ c !! (m + size b)))" apply (unfold word_split_bin' test_bit_bin) apply (clarify) apply (clarsimp simp: word_ubin.eq_norm nth_bintr word_size split: prod.splits) apply (drule bin_nth_split) apply safe apply (simp_all add: add.commute) apply (erule bin_nth_uint_imp)+ done lemma test_bit_split: "word_split c = (a, b) \ (\n::nat. b !! n \ n < size b \ c !! n) \ (\m::nat. a !! m \ m < size a \ c !! (m + size b))" by (simp add: test_bit_split') lemma test_bit_split_eq: "word_split c = (a, b) \ ((\n::nat. b !! n = (n < size b \ c !! n)) \ (\m::nat. a !! m = (m < size a \ c !! (m + size b))))" apply (rule_tac iffI) apply (rule_tac conjI) apply (erule test_bit_split [THEN conjunct1]) apply (erule test_bit_split [THEN conjunct2]) apply (case_tac "word_split c") apply (frule test_bit_split) apply (erule trans) apply (fastforce intro!: word_eqI simp add: word_size) done \ \this odd result is analogous to \ucast_id\, result to the length given by the result type\ lemma word_cat_id: "word_cat a b = b" by (simp add: word_cat_bin' word_ubin.inverse_norm) \ \limited hom result\ lemma word_cat_hom: "LENGTH('a::len0) \ LENGTH('b::len0) + LENGTH('c::len0) \ (word_cat (word_of_int w :: 'b word) (b :: 'c word) :: 'a word) = word_of_int (bin_cat w (size b) (uint b))" by (auto simp: word_cat_def word_size word_ubin.norm_eq_iff [symmetric] word_ubin.eq_norm bintr_cat min.absorb1) lemma word_cat_split_alt: "size w \ size u + size v \ word_split w = (u, v) \ word_cat u v = w" apply (rule word_eqI) apply (drule test_bit_split) apply (clarsimp simp add : test_bit_cat word_size) apply safe apply arith done lemmas word_cat_split_size = sym [THEN [2] word_cat_split_alt [symmetric]] subsubsection \Split and slice\ lemma split_slices: "word_split w = (u, v) \ u = slice (size v) w \ v = slice 0 w" apply (drule test_bit_split) apply (rule conjI) apply (rule word_eqI, clarsimp simp: nth_slice word_size)+ done lemma slice_cat1 [OF refl]: "wc = word_cat a b \ size wc >= size a + size b \ slice (size b) wc = a" apply safe apply (rule word_eqI) apply (simp add: nth_slice test_bit_cat word_size) done lemmas slice_cat2 = trans [OF slice_id word_cat_id] lemma cat_slices: "a = slice n c \ b = slice 0 c \ n = size b \ size a + size b >= size c \ word_cat a b = c" apply safe apply (rule word_eqI) apply (simp add: nth_slice test_bit_cat word_size) apply safe apply arith done lemma word_split_cat_alt: "w = word_cat u v \ size u + size v \ size w \ word_split w = (u, v)" apply (case_tac "word_split w") apply (rule trans, assumption) apply (drule test_bit_split) apply safe apply (rule word_eqI, clarsimp simp: test_bit_cat word_size)+ done lemmas word_cat_bl_no_bin [simp] = word_cat_bl [where a="numeral a" and b="numeral b", unfolded to_bl_numeral] for a b (* FIXME: negative numerals, 0 and 1 *) lemmas word_split_bl_no_bin [simp] = word_split_bl_eq [where c="numeral c", unfolded to_bl_numeral] for c text \ This odd result arises from the fact that the statement of the result implies that the decoded words are of the same type, and therefore of the same length, as the original word.\ lemma word_rsplit_same: "word_rsplit w = [w]" by (simp add: word_rsplit_def bin_rsplit_all) lemma word_rsplit_empty_iff_size: "word_rsplit w = [] \ size w = 0" by (simp add: word_rsplit_def bin_rsplit_def word_size bin_rsplit_aux_simp_alt Let_def split: prod.split) lemma test_bit_rsplit: "sw = word_rsplit w \ m < size (hd sw) \ k < length sw \ (rev sw ! k) !! m = w !! (k * size (hd sw) + m)" for sw :: "'a::len word list" apply (unfold word_rsplit_def word_test_bit_def) apply (rule trans) apply (rule_tac f = "\x. bin_nth x m" in arg_cong) apply (rule nth_map [symmetric]) apply simp apply (rule bin_nth_rsplit) apply simp_all apply (simp add : word_size rev_map) apply (rule trans) defer apply (rule map_ident [THEN fun_cong]) apply (rule refl [THEN map_cong]) apply (simp add : word_ubin.eq_norm) apply (erule bin_rsplit_size_sign [OF len_gt_0 refl]) done lemma word_rcat_bl: "word_rcat wl = of_bl (concat (map to_bl wl))" by (auto simp: word_rcat_def to_bl_def' of_bl_def bin_rcat_bl) lemma size_rcat_lem': "size (concat (map to_bl wl)) = length wl * size (hd wl)" by (induct wl) (auto simp: word_size) lemmas size_rcat_lem = size_rcat_lem' [unfolded word_size] lemmas td_gal_lt_len = len_gt_0 [THEN td_gal_lt] lemma nth_rcat_lem: "n < length (wl::'a word list) * LENGTH('a::len) \ rev (concat (map to_bl wl)) ! n = rev (to_bl (rev wl ! (n div LENGTH('a)))) ! (n mod LENGTH('a))" apply (induct wl) apply clarsimp apply (clarsimp simp add : nth_append size_rcat_lem) apply (simp (no_asm_use) only: mult_Suc [symmetric] td_gal_lt_len less_Suc_eq_le minus_div_mult_eq_mod [symmetric]) apply clarsimp done lemma test_bit_rcat: "sw = size (hd wl) \ rc = word_rcat wl \ rc !! n = (n < size rc \ n div sw < size wl \ (rev wl) ! (n div sw) !! (n mod sw))" for wl :: "'a::len word list" apply (unfold word_rcat_bl word_size) apply (clarsimp simp add: test_bit_of_bl size_rcat_lem word_size td_gal_lt_len) apply safe apply (auto simp: test_bit_bl word_size td_gal_lt_len [THEN iffD2, THEN nth_rcat_lem]) done lemma foldl_eq_foldr: "foldl (+) x xs = foldr (+) (x # xs) 0" for x :: "'a::comm_monoid_add" by (induct xs arbitrary: x) (auto simp: add.assoc) lemmas test_bit_cong = arg_cong [where f = "test_bit", THEN fun_cong] lemmas test_bit_rsplit_alt = trans [OF nth_rev_alt [THEN test_bit_cong] test_bit_rsplit [OF refl asm_rl diff_Suc_less]] \ \lazy way of expressing that u and v, and su and sv, have same types\ lemma word_rsplit_len_indep [OF refl refl refl refl]: "[u,v] = p \ [su,sv] = q \ word_rsplit u = su \ word_rsplit v = sv \ length su = length sv" by (auto simp: word_rsplit_def bin_rsplit_len_indep) lemma length_word_rsplit_size: "n = LENGTH('a::len) \ length (word_rsplit w :: 'a word list) \ m \ size w \ m * n" by (auto simp: word_rsplit_def word_size bin_rsplit_len_le) lemmas length_word_rsplit_lt_size = length_word_rsplit_size [unfolded Not_eq_iff linorder_not_less [symmetric]] lemma length_word_rsplit_exp_size: "n = LENGTH('a::len) \ length (word_rsplit w :: 'a word list) = (size w + n - 1) div n" by (auto simp: word_rsplit_def word_size bin_rsplit_len) lemma length_word_rsplit_even_size: "n = LENGTH('a::len) \ size w = m * n \ length (word_rsplit w :: 'a word list) = m" by (auto simp: length_word_rsplit_exp_size given_quot_alt) lemmas length_word_rsplit_exp_size' = refl [THEN length_word_rsplit_exp_size] \ \alternative proof of \word_rcat_rsplit\\ lemmas tdle = times_div_less_eq_dividend lemmas dtle = xtr4 [OF tdle mult.commute] lemma word_rcat_rsplit: "word_rcat (word_rsplit w) = w" apply (rule word_eqI) apply (clarsimp simp: test_bit_rcat word_size) apply (subst refl [THEN test_bit_rsplit]) apply (simp_all add: word_size refl [THEN length_word_rsplit_size [simplified not_less [symmetric], simplified]]) apply safe apply (erule xtr7, rule dtle)+ done lemma size_word_rsplit_rcat_size: "word_rcat ws = frcw \ size frcw = length ws * LENGTH('a) \ length (word_rsplit frcw::'a word list) = length ws" for ws :: "'a::len word list" and frcw :: "'b::len0 word" apply (clarsimp simp: word_size length_word_rsplit_exp_size') apply (fast intro: given_quot_alt) done lemma msrevs: "0 < n \ (k * n + m) div n = m div n + k" "(k * n + m) mod n = m mod n" for n :: nat by (auto simp: add.commute) lemma word_rsplit_rcat_size [OF refl]: "word_rcat ws = frcw \ size frcw = length ws * LENGTH('a) \ word_rsplit frcw = ws" for ws :: "'a::len word list" apply (frule size_word_rsplit_rcat_size, assumption) apply (clarsimp simp add : word_size) apply (rule nth_equalityI, assumption) apply clarsimp apply (rule word_eqI [rule_format]) apply (rule trans) apply (rule test_bit_rsplit_alt) apply (clarsimp simp: word_size)+ apply (rule trans) apply (rule test_bit_rcat [OF refl refl]) apply (simp add: word_size) apply (subst nth_rev) apply arith apply (simp add: le0 [THEN [2] xtr7, THEN diff_Suc_less]) apply safe apply (simp add: diff_mult_distrib) apply (rule mpl_lem) apply (cases "size ws") apply simp_all done subsection \Rotation\ lemmas rotater_0' [simp] = rotater_def [where n = "0", simplified] lemmas word_rot_defs = word_roti_def word_rotr_def word_rotl_def 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 lemmas rotate_eqs = trans [OF rotate0 [THEN fun_cong] id_apply] rotate_rotate [symmetric] rotate_id rotate_conv_mod rotate_eq_mod subsubsection \Rotation of list to right\ 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 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 restrict_to_left: "x = y \ x = z \ y = z" by simp 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) subsubsection \map, map2, commuting with rotate(r)\ 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) \ \corresponding equalities for word rotation\ lemma to_bl_rotl: "to_bl (word_rotl n w) = rotate n (to_bl w)" by (simp add: word_bl.Abs_inverse' word_rotl_def) lemmas blrs0 = rotate_eqs [THEN to_bl_rotl [THEN trans]] lemmas word_rotl_eqs = blrs0 [simplified word_bl_Rep' word_bl.Rep_inject to_bl_rotl [symmetric]] lemma to_bl_rotr: "to_bl (word_rotr n w) = rotater n (to_bl w)" by (simp add: word_bl.Abs_inverse' word_rotr_def) lemmas brrs0 = rotater_eqs [THEN to_bl_rotr [THEN trans]] lemmas word_rotr_eqs = brrs0 [simplified word_bl_Rep' word_bl.Rep_inject to_bl_rotr [symmetric]] declare word_rotr_eqs (1) [simp] declare word_rotl_eqs (1) [simp] lemma word_rot_rl [simp]: "word_rotl k (word_rotr k v) = v" and word_rot_lr [simp]: "word_rotr k (word_rotl k v) = v" by (auto simp add: to_bl_rotr to_bl_rotl word_bl.Rep_inject [symmetric]) lemma word_rot_gal: "word_rotr n v = w \ word_rotl n w = v" and word_rot_gal': "w = word_rotr n v \ v = word_rotl n w" by (auto simp: to_bl_rotr to_bl_rotl word_bl.Rep_inject [symmetric] dest: sym) lemma word_rotr_rev: "word_rotr n w = word_reverse (word_rotl n (word_reverse w))" by (simp only: word_bl.Rep_inject [symmetric] to_bl_word_rev to_bl_rotr to_bl_rotl rotater_rev) lemma word_roti_0 [simp]: "word_roti 0 w = w" by (auto simp: word_rot_defs) lemmas abl_cong = arg_cong [where f = "of_bl"] lemma word_roti_add: "word_roti (m + n) w = word_roti m (word_roti n w)" proof - have rotater_eq_lem: "\m n xs. m = n \ rotater m xs = rotater n xs" by auto have rotate_eq_lem: "\m n xs. m = n \ rotate m xs = rotate n xs" by auto note rpts [symmetric] = rotate_inv_plus [THEN conjunct1] rotate_inv_plus [THEN conjunct2, THEN conjunct1] rotate_inv_plus [THEN conjunct2, THEN conjunct2, THEN conjunct1] rotate_inv_plus [THEN conjunct2, THEN conjunct2, THEN conjunct2] note rrp = trans [symmetric, OF rotate_rotate rotate_eq_lem] note rrrp = trans [symmetric, OF rotater_add [symmetric] rotater_eq_lem] show ?thesis apply (unfold word_rot_defs) apply (simp only: split: if_split) apply (safe intro!: abl_cong) apply (simp_all only: to_bl_rotl [THEN word_bl.Rep_inverse'] to_bl_rotl to_bl_rotr [THEN word_bl.Rep_inverse'] to_bl_rotr) apply (rule rrp rrrp rpts, simp add: nat_add_distrib [symmetric] nat_diff_distrib [symmetric])+ done qed lemma word_roti_conv_mod': "word_roti n w = word_roti (n mod int (size w)) w" proof (cases "size w = 0") case True then show ?thesis by simp next case False then have [simp]: "l mod int (size w) \ 0" for l by simp then have *: "word_roti (n mod int (size w)) w = word_rotr (nat (n mod int (size w))) w" by (simp add: word_roti_def) show ?thesis proof (cases "n \ 0") case True then show ?thesis apply (auto simp add: not_le *) apply (auto simp add: word_rot_defs) apply (safe intro!: abl_cong) apply (rule rotater_eqs) apply (simp add: word_size nat_mod_distrib) done next case False moreover define k where "k = - n" ultimately have n: "n = - k" by simp_all from False show ?thesis apply (auto simp add: not_le *) apply (auto simp add: word_rot_defs) apply (simp add: n) apply (safe intro!: abl_cong) apply (simp add: rotater_add [symmetric] rotate_gal [symmetric]) apply (rule rotater_eqs) apply (simp add: word_size [symmetric, of w]) apply (rule of_nat_eq_0_iff [THEN iffD1]) apply (auto simp add: nat_add_distrib [symmetric] mod_eq_0_iff_dvd) using dvd_nat_abs_iff [of "size w" "- k mod int (size w) + k"] apply (simp add: algebra_simps) apply (simp add: word_size) - apply (metis (no_types, hide_lams) add.right_inverse dvd_0_right dvd_mod_imp_dvd dvd_refl minus_dvd_iff minus_equation_iff mod_0 mod_add_eq mod_minus_eq) + apply (metis dvd_eq_mod_eq_0 eq_neg_iff_add_eq_0 k_def mod_0 mod_add_right_eq n) done qed qed lemmas word_roti_conv_mod = word_roti_conv_mod' [unfolded word_size] subsubsection \"Word rotation commutes with bit-wise operations\ \ \using locale to not pollute lemma namespace\ locale word_rotate begin lemmas word_rot_defs' = to_bl_rotl to_bl_rotr lemmas blwl_syms [symmetric] = bl_word_not bl_word_and bl_word_or bl_word_xor lemmas lbl_lbl = trans [OF word_bl_Rep' word_bl_Rep' [symmetric]] lemmas ths_map2 [OF lbl_lbl] = rotate_map2 rotater_map2 lemmas ths_map [where xs = "to_bl v"] = rotate_map rotater_map for v lemmas th1s [simplified word_rot_defs' [symmetric]] = ths_map2 ths_map lemma word_rot_logs: "word_rotl n (NOT v) = NOT (word_rotl n v)" "word_rotr n (NOT v) = NOT (word_rotr n v)" "word_rotl n (x AND y) = word_rotl n x AND word_rotl n y" "word_rotr n (x AND y) = word_rotr n x AND word_rotr n y" "word_rotl n (x OR y) = word_rotl n x OR word_rotl n y" "word_rotr n (x OR y) = word_rotr n x OR word_rotr n y" "word_rotl n (x XOR y) = word_rotl n x XOR word_rotl n y" "word_rotr n (x XOR y) = word_rotr n x XOR word_rotr n y" by (rule word_bl.Rep_eqD, rule word_rot_defs' [THEN trans], simp only: blwl_syms [symmetric], rule th1s [THEN trans], rule refl)+ end lemmas word_rot_logs = word_rotate.word_rot_logs lemmas bl_word_rotl_dt = trans [OF to_bl_rotl rotate_drop_take, simplified word_bl_Rep'] lemmas bl_word_rotr_dt = trans [OF to_bl_rotr rotater_drop_take, simplified word_bl_Rep'] lemma bl_word_roti_dt': "n = nat ((- i) mod int (size (w :: 'a::len word))) \ to_bl (word_roti i w) = drop n (to_bl w) @ take n (to_bl w)" apply (unfold word_roti_def) apply (simp add: bl_word_rotl_dt bl_word_rotr_dt word_size) apply safe apply (simp add: zmod_zminus1_eq_if) apply safe apply (simp add: nat_mult_distrib) apply (simp add: nat_diff_distrib [OF pos_mod_sign pos_mod_conj [THEN conjunct2, THEN order_less_imp_le]] nat_mod_distrib) apply (simp add: nat_mod_distrib) done lemmas bl_word_roti_dt = bl_word_roti_dt' [unfolded word_size] lemmas word_rotl_dt = bl_word_rotl_dt [THEN word_bl.Rep_inverse' [symmetric]] lemmas word_rotr_dt = bl_word_rotr_dt [THEN word_bl.Rep_inverse' [symmetric]] lemmas word_roti_dt = bl_word_roti_dt [THEN word_bl.Rep_inverse' [symmetric]] lemma word_rotx_0 [simp] : "word_rotr i 0 = 0 \ word_rotl i 0 = 0" by (simp add: word_rotr_dt word_rotl_dt replicate_add [symmetric]) lemma word_roti_0' [simp] : "word_roti n 0 = 0" by (auto simp: word_roti_def) lemmas word_rotr_dt_no_bin' [simp] = word_rotr_dt [where w="numeral w", unfolded to_bl_numeral] for w (* FIXME: negative numerals, 0 and 1 *) lemmas word_rotl_dt_no_bin' [simp] = word_rotl_dt [where w="numeral w", unfolded to_bl_numeral] for w (* FIXME: negative numerals, 0 and 1 *) declare word_roti_def [simp] subsection \Maximum machine word\ lemma word_int_cases: fixes x :: "'a::len0 word" obtains n where "x = word_of_int n" and "0 \ n" and "n < 2^LENGTH('a)" by (cases x rule: word_uint.Abs_cases) (simp add: uints_num) lemma word_nat_cases [cases type: word]: fixes x :: "'a::len word" obtains n where "x = of_nat n" and "n < 2^LENGTH('a)" by (cases x rule: word_unat.Abs_cases) (simp add: unats_def) lemma max_word_eq: "(max_word::'a::len word) = 2^LENGTH('a) - 1" by (simp add: max_word_def word_of_int_hom_syms word_of_int_2p) lemma max_word_max [simp,intro!]: "n \ max_word" by (cases n rule: word_int_cases) - (simp add: max_word_def word_le_def int_word_uint mod_pos_pos_trivial del: minus_mod_self1) + (simp add: max_word_def word_le_def int_word_uint del: minus_mod_self1) lemma word_of_int_2p_len: "word_of_int (2 ^ LENGTH('a)) = (0::'a::len0 word)" by (subst word_uint.Abs_norm [symmetric]) simp lemma word_pow_0: "(2::'a::len word) ^ LENGTH('a) = 0" proof - have "word_of_int (2 ^ LENGTH('a)) = (0::'a word)" by (rule word_of_int_2p_len) then show ?thesis by (simp add: word_of_int_2p) qed lemma max_word_wrap: "x + 1 = 0 \ x = max_word" apply (simp add: max_word_eq) apply uint_arith apply (auto simp: word_pow_0) done lemma max_word_minus: "max_word = (-1::'a::len word)" proof - have "-1 + 1 = (0::'a word)" by simp then show ?thesis by (rule max_word_wrap [symmetric]) qed lemma max_word_bl [simp]: "to_bl (max_word::'a::len word) = replicate (LENGTH('a)) True" by (subst max_word_minus to_bl_n1)+ simp lemma max_test_bit [simp]: "(max_word::'a::len word) !! n \ n < LENGTH('a)" by (auto simp: test_bit_bl word_size) lemma word_and_max [simp]: "x AND max_word = x" by (rule word_eqI) (simp add: word_ops_nth_size word_size) lemma word_or_max [simp]: "x OR max_word = max_word" by (rule word_eqI) (simp add: word_ops_nth_size word_size) lemma word_ao_dist2: "x AND (y OR z) = x AND y OR x AND z" for x y z :: "'a::len0 word" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) lemma word_oa_dist2: "x OR y AND z = (x OR y) AND (x OR z)" for x y z :: "'a::len0 word" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) lemma word_and_not [simp]: "x AND NOT x = 0" for x :: "'a::len0 word" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) lemma word_or_not [simp]: "x OR NOT x = max_word" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) global_interpretation word_bool_alg: boolean_algebra where conj = "(AND)" and disj = "(OR)" and compl = NOT and zero = 0 and one = max_word rewrites "word_bool_alg.xor = (XOR)" proof - interpret boolean_algebra where conj = "(AND)" and disj = "(OR)" and compl = NOT and zero = 0 and one = max_word apply standard apply (simp_all add: word_bw_assocs word_bw_comms word_bw_lcs) apply (fact word_ao_dist2) apply (fact word_oa_dist2) done show "boolean_algebra (AND) (OR) NOT 0 max_word" .. show "xor = (XOR)" by (auto simp add: fun_eq_iff word_eq_iff xor_def word_ops_nth_size word_size) qed lemma word_xor_and_or: "x XOR y = x AND NOT y OR NOT x AND y" for x y :: "'a::len0 word" by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) lemma shiftr_x_0 [iff]: "x >> 0 = x" for x :: "'a::len0 word" by (simp add: shiftr_bl) lemma shiftl_x_0 [simp]: "x << 0 = x" for x :: "'a::len word" by (simp add: shiftl_t2n) lemma shiftl_1 [simp]: "(1::'a::len word) << n = 2^n" by (simp add: shiftl_t2n) lemma uint_lt_0 [simp]: "uint x < 0 = False" by (simp add: linorder_not_less) lemma shiftr1_1 [simp]: "shiftr1 (1::'a::len word) = 0" unfolding shiftr1_def by simp lemma shiftr_1[simp]: "(1::'a::len word) >> n = (if n = 0 then 1 else 0)" by (induct n) (auto simp: shiftr_def) lemma word_less_1 [simp]: "x < 1 \ x = 0" for x :: "'a::len word" by (simp add: word_less_nat_alt unat_0_iff) lemma to_bl_mask: "to_bl (mask n :: 'a::len word) = replicate (LENGTH('a) - n) False @ replicate (min (LENGTH('a)) n) True" by (simp add: mask_bl word_rep_drop min_def) lemma map_replicate_True: "n = length xs \ map (\(x,y). x \ y) (zip xs (replicate n True)) = xs" by (induct xs arbitrary: n) auto lemma map_replicate_False: "n = length xs \ map (\(x,y). x \ y) (zip xs (replicate n False)) = replicate n False" by (induct xs arbitrary: n) auto lemma bl_and_mask: fixes w :: "'a::len word" and n :: nat defines "n' \ LENGTH('a) - n" shows "to_bl (w AND mask n) = replicate n' False @ drop n' (to_bl w)" proof - note [simp] = map_replicate_True map_replicate_False have "to_bl (w AND mask n) = map2 (\) (to_bl w) (to_bl (mask n::'a::len word))" by (simp add: bl_word_and) also have "to_bl w = take n' (to_bl w) @ drop n' (to_bl w)" by simp also have "map2 (\) \ (to_bl (mask n::'a::len word)) = replicate n' False @ drop n' (to_bl w)" unfolding to_bl_mask n'_def by (subst zip_append) auto finally show ?thesis . qed lemma drop_rev_takefill: "length xs \ n \ drop (n - length xs) (rev (takefill False n (rev xs))) = xs" by (simp add: takefill_alt rev_take) lemma map_nth_0 [simp]: "map ((!!) (0::'a::len0 word)) xs = replicate (length xs) False" by (induct xs) auto lemma uint_plus_if_size: "uint (x + y) = (if uint x + uint y < 2^size x then uint x + uint y else uint x + uint y - 2^size x)" by (simp add: word_arith_wis int_word_uint mod_add_if_z word_size) lemma unat_plus_if_size: "unat (x + y) = (if unat x + unat y < 2^size x then unat x + unat y else unat x + unat y - 2^size x)" for x y :: "'a::len word" apply (subst word_arith_nat_defs) apply (subst unat_of_nat) apply (simp add: mod_nat_add word_size) done lemma word_neq_0_conv: "w \ 0 \ 0 < w" for w :: "'a::len word" by (simp add: word_gt_0) lemma max_lt: "unat (max a b div c) = unat (max a b) div unat c" for c :: "'a::len word" by (fact unat_div) lemma uint_sub_if_size: "uint (x - y) = (if uint y \ uint x then uint x - uint y else uint x - uint y + 2^size x)" by (simp add: word_arith_wis int_word_uint mod_sub_if_z word_size) lemma unat_sub: "b \ a \ unat (a - b) = unat a - unat b" by (simp add: unat_def uint_sub_if_size word_le_def nat_diff_distrib) lemmas word_less_sub1_numberof [simp] = word_less_sub1 [of "numeral w"] for w lemmas word_le_sub1_numberof [simp] = word_le_sub1 [of "numeral w"] for w lemma word_of_int_minus: "word_of_int (2^LENGTH('a) - i) = (word_of_int (-i)::'a::len word)" proof - have *: "2^LENGTH('a) - i = -i + 2^LENGTH('a)" by simp show ?thesis apply (subst *) apply (subst word_uint.Abs_norm [symmetric], subst mod_add_self2) apply simp done qed lemmas word_of_int_inj = word_uint.Abs_inject [unfolded uints_num, simplified] lemma word_le_less_eq: "x \ y \ x = y \ x < y" for x y :: "'z::len word" by (auto simp add: order_class.le_less) lemma mod_plus_cong: fixes b b' :: int assumes 1: "b = b'" and 2: "x mod b' = x' mod b'" and 3: "y mod b' = y' mod b'" and 4: "x' + y' = z'" shows "(x + y) mod b = z' mod b'" proof - from 1 2[symmetric] 3[symmetric] have "(x + y) mod b = (x' mod b' + y' mod b') mod b'" by (simp add: mod_add_eq) also have "\ = (x' + y') mod b'" by (simp add: mod_add_eq) finally show ?thesis by (simp add: 4) qed lemma mod_minus_cong: fixes b b' :: int assumes "b = b'" and "x mod b' = x' mod b'" and "y mod b' = y' mod b'" and "x' - y' = z'" shows "(x - y) mod b = z' mod b'" using assms [symmetric] by (auto intro: mod_diff_cong) lemma word_induct_less: "P 0 \ (\n. n < m \ P n \ P (1 + n)) \ P m" for P :: "'a::len word \ bool" apply (cases m) apply atomize apply (erule rev_mp)+ apply (rule_tac x=m in spec) apply (induct_tac n) apply simp apply clarsimp apply (erule impE) apply clarsimp apply (erule_tac x=n in allE) apply (erule impE) apply (simp add: unat_arith_simps) apply (clarsimp simp: unat_of_nat) apply simp apply (erule_tac x="of_nat na" in allE) apply (erule impE) apply (simp add: unat_arith_simps) apply (clarsimp simp: unat_of_nat) apply simp done lemma word_induct: "P 0 \ (\n. P n \ P (1 + n)) \ P m" for P :: "'a::len word \ bool" by (erule word_induct_less) simp lemma word_induct2 [induct type]: "P 0 \ (\n. 1 + n \ 0 \ P n \ P (1 + n)) \ P n" for P :: "'b::len word \ bool" apply (rule word_induct) apply simp apply (case_tac "1 + n = 0") apply auto done subsection \Recursion combinator for words\ definition word_rec :: "'a \ ('b::len word \ 'a \ 'a) \ 'b word \ 'a" where "word_rec forZero forSuc n = rec_nat forZero (forSuc \ of_nat) (unat n)" lemma word_rec_0: "word_rec z s 0 = z" by (simp add: word_rec_def) lemma word_rec_Suc: "1 + n \ 0 \ word_rec z s (1 + n) = s n (word_rec z s n)" for n :: "'a::len word" apply (simp add: word_rec_def unat_word_ariths) apply (subst nat_mod_eq') apply (metis Suc_eq_plus1_left Suc_lessI of_nat_2p unat_1 unat_lt2p word_arith_nat_add) apply simp done lemma word_rec_Pred: "n \ 0 \ word_rec z s n = s (n - 1) (word_rec z s (n - 1))" apply (rule subst[where t="n" and s="1 + (n - 1)"]) apply simp apply (subst word_rec_Suc) apply simp apply simp done lemma word_rec_in: "f (word_rec z (\_. f) n) = word_rec (f z) (\_. f) n" by (induct n) (simp_all add: word_rec_0 word_rec_Suc) lemma word_rec_in2: "f n (word_rec z f n) = word_rec (f 0 z) (f \ (+) 1) n" by (induct n) (simp_all add: word_rec_0 word_rec_Suc) lemma word_rec_twice: "m \ n \ word_rec z f n = word_rec (word_rec z f (n - m)) (f \ (+) (n - m)) m" apply (erule rev_mp) apply (rule_tac x=z in spec) apply (rule_tac x=f in spec) apply (induct n) apply (simp add: word_rec_0) apply clarsimp apply (rule_tac t="1 + n - m" and s="1 + (n - m)" in subst) apply simp apply (case_tac "1 + (n - m) = 0") apply (simp add: word_rec_0) apply (rule_tac f = "word_rec a b" for a b in arg_cong) apply (rule_tac t="m" and s="m + (1 + (n - m))" in subst) apply simp apply (simp (no_asm_use)) apply (simp add: word_rec_Suc word_rec_in2) apply (erule impE) apply uint_arith apply (drule_tac x="x \ (+) 1" in spec) apply (drule_tac x="x 0 xa" in spec) apply simp apply (rule_tac t="\a. x (1 + (n - m + a))" and s="\a. x (1 + (n - m) + a)" in subst) apply (clarsimp simp add: fun_eq_iff) apply (rule_tac t="(1 + (n - m + xb))" and s="1 + (n - m) + xb" in subst) apply simp apply (rule refl) apply (rule refl) done lemma word_rec_id: "word_rec z (\_. id) n = z" by (induct n) (auto simp add: word_rec_0 word_rec_Suc) lemma word_rec_id_eq: "\m < n. f m = id \ word_rec z f n = z" apply (erule rev_mp) apply (induct n) apply (auto simp add: word_rec_0 word_rec_Suc) apply (drule spec, erule mp) apply uint_arith apply (drule_tac x=n in spec, erule impE) apply uint_arith apply simp done lemma word_rec_max: "\m\n. m \ - 1 \ f m = id \ word_rec z f (- 1) = word_rec z f n" apply (subst word_rec_twice[where n="-1" and m="-1 - n"]) apply simp apply simp apply (rule word_rec_id_eq) apply clarsimp apply (drule spec, rule mp, erule mp) apply (rule word_plus_mono_right2[OF _ order_less_imp_le]) prefer 2 apply assumption apply simp apply (erule contrapos_pn) apply simp apply (drule arg_cong[where f="\x. x - n"]) apply simp done subsection \More\ lemma test_bit_1' [simp]: "(1 :: 'a :: len0 word) !! n \ 0 < LENGTH('a) \ n = 0" by (cases n) (simp_all only: one_word_def test_bit_wi bin_nth.simps, simp_all) lemma mask_0 [simp]: "mask 0 = 0" by (simp add: Word.mask_def) lemma shiftl0 [simp]: "x << 0 = (x :: 'a :: len0 word)" by (metis shiftl_rev shiftr_x_0 word_rev_gal) lemma mask_1: "mask 1 = 1" by (simp add: mask_def) lemma mask_Suc_0: "mask (Suc 0) = 1" by (simp add: mask_def) lemma mask_numeral: "mask (numeral n) = 2 * mask (pred_numeral n) + 1" by (simp add: mask_def neg_numeral_class.sub_def numeral_eq_Suc numeral_pow) lemma bin_last_bintrunc: "bin_last (bintrunc l n) = (l > 0 \ bin_last n)" by (cases l) simp_all lemma word_and_1: "n AND 1 = (if n !! 0 then 1 else 0)" for n :: "_ word" by transfer (rule bin_rl_eqI, simp_all add: bin_rest_trunc bin_last_bintrunc) lemma bintrunc_shiftl: "bintrunc n (m << i) = bintrunc (n - i) m << i" proof (induction i arbitrary: n) case 0 show ?case by simp next case (Suc i) then show ?case by (cases n) simp_all qed lemma shiftl_transfer [transfer_rule]: includes lifting_syntax shows "(pcr_word ===> (=) ===> pcr_word) (<<) (<<)" by (auto intro!: rel_funI word_eqI simp add: word.pcr_cr_eq cr_word_def word_size nth_shiftl) lemma uint_shiftl: "uint (n << i) = bintrunc (size n) (uint n << i)" unfolding word_size by transfer (simp add: bintrunc_shiftl) subsection \Misc\ declare bin_to_bl_def [simp] ML_file \Tools/word_lib.ML\ ML_file \Tools/smt_word.ML\ hide_const (open) Word end