# HG changeset patch # User haftmann # Date 1571510469 -7200 # Node ID 8c2bef3df48885f0cf3e2ba73579f1337690b84d # Parent 38298c04c12e4f373fbe2a649306686fad3a7dff refined proof of concept for bit operations diff -r 38298c04c12e -r 8c2bef3df488 src/HOL/ex/Bit_Lists.thy --- a/src/HOL/ex/Bit_Lists.thy Sat Oct 19 20:41:03 2019 +0200 +++ b/src/HOL/ex/Bit_Lists.thy Sat Oct 19 20:41:09 2019 +0200 @@ -1,169 +1,938 @@ (* Author: Florian Haftmann, TUM *) -section \Proof of concept for algebraically founded lists of bits\ +section \Proof(s) of concept for algebraically founded lists of bits\ theory Bit_Lists imports Main begin +subsection \Bit representations\ + +subsubsection \Mere syntactic bit representation\ + +class bit_representation = + fixes bits_of :: "'a \ bool list" + and of_bits :: "bool list \ 'a" + assumes of_bits_of [simp]: "of_bits (bits_of a) = a" + + +subsubsection \Algebraic bit representation\ + context comm_semiring_1 begin + +primrec radix_value :: "('b \ 'a) \ 'a \ 'b list \ 'a" + where "radix_value f b [] = 0" + | "radix_value f b (a # as) = f a + radix_value f b as * b" + +abbreviation (input) unsigned_of_bits :: "bool list \ 'a" + where "unsigned_of_bits \ radix_value of_bool 2" + +lemma unsigned_of_bits_replicate_False [simp]: + "unsigned_of_bits (replicate n False) = 0" + by (induction n) simp_all + +end + +context unique_euclidean_semiring_with_nat +begin -primrec of_unsigned :: "bool list \ 'a" - where "of_unsigned [] = 0" - | "of_unsigned (b # bs) = of_bool b + 2 * of_unsigned bs" +lemma unsigned_of_bits_append [simp]: + "unsigned_of_bits (bs @ cs) = unsigned_of_bits bs + + push_bit (length bs) (unsigned_of_bits cs)" + by (induction bs) (simp_all add: push_bit_double, + simp_all add: algebra_simps) + +lemma unsigned_of_bits_take [simp]: + "unsigned_of_bits (take n bs) = take_bit n (unsigned_of_bits bs)" +proof (induction bs arbitrary: n) + case Nil + then show ?case + by simp +next + case (Cons b bs) + then show ?case + by (cases n) (simp_all add: ac_simps) +qed + +lemma unsigned_of_bits_drop [simp]: + "unsigned_of_bits (drop n bs) = drop_bit n (unsigned_of_bits bs)" +proof (induction bs arbitrary: n) + case Nil + then show ?case + by simp +next + case (Cons b bs) + then show ?case + by (cases n) simp_all +qed + +end + + +subsubsection \Instances\ + +text \Unclear whether a \<^typ>\bool\ instantiation is needed or not\ + +instantiation nat :: bit_representation +begin + +fun bits_of_nat :: "nat \ bool list" + where "bits_of (n::nat) = + (if n = 0 then [] else odd n # bits_of (n div 2))" + +lemma bits_of_nat_simps [simp]: + "bits_of (0::nat) = []" + "n > 0 \ bits_of n = odd n # bits_of (n div 2)" for n :: nat + by simp_all + +declare bits_of_nat.simps [simp del] + +definition of_bits_nat :: "bool list \ nat" + where [simp]: "of_bits_nat = unsigned_of_bits" + \ \remove simp\ + +instance proof + show "of_bits (bits_of n) = n" for n :: nat + by (induction n rule: nat_bit_induct) simp_all +qed end -context comm_ring_1 +lemma bits_of_Suc_0 [simp]: + "bits_of (Suc 0) = [True]" + by simp + +lemma bits_of_1_nat [simp]: + "bits_of (1 :: nat) = [True]" + by simp + +lemma bits_of_nat_numeral_simps [simp]: + "bits_of (numeral Num.One :: nat) = [True]" (is ?One) + "bits_of (numeral (Num.Bit0 n) :: nat) = False # bits_of (numeral n :: nat)" (is ?Bit0) + "bits_of (numeral (Num.Bit1 n) :: nat) = True # bits_of (numeral n :: nat)" (is ?Bit1) +proof - + show ?One + by simp + define m :: nat where "m = numeral n" + then have "m > 0" and *: "numeral n = m" "numeral (Num.Bit0 n) = 2 * m" "numeral (Num.Bit1 n) = Suc (2 * m)" + by simp_all + from \m > 0\ show ?Bit0 ?Bit1 + by (simp_all add: *) +qed + +lemma unsigned_of_bits_of_nat [simp]: + "unsigned_of_bits (bits_of n) = n" for n :: nat + using of_bits_of [of n] by simp + +instantiation int :: bit_representation begin -definition of_signed :: "bool list \ 'a" - where "of_signed bs = (if bs = [] then 0 else if last bs - then - (of_unsigned (map Not bs) + 1) else of_unsigned bs)" +fun bits_of_int :: "int \ bool list" + where "bits_of_int k = odd k # + (if k = 0 \ k = - 1 then [] else bits_of_int (k div 2))" + +lemma bits_of_int_simps [simp]: + "bits_of (0 :: int) = [False]" + "bits_of (- 1 :: int) = [True]" + "k \ 0 \ k \ - 1 \ bits_of k = odd k # bits_of (k div 2)" for k :: int + by simp_all + +lemma bits_of_not_Nil [simp]: + "bits_of k \ []" for k :: int + by simp + +declare bits_of_int.simps [simp del] + +definition of_bits_int :: "bool list \ int" + where "of_bits_int bs = (if bs = [] \ \ last bs then unsigned_of_bits bs + else unsigned_of_bits bs - 2 ^ length bs)" + +lemma of_bits_int_simps [simp]: + "of_bits [] = (0 :: int)" + "of_bits [False] = (0 :: int)" + "of_bits [True] = (- 1 :: int)" + "of_bits (bs @ [b]) = (unsigned_of_bits bs :: int) - (2 ^ length bs) * of_bool b" + "of_bits (False # bs) = 2 * (of_bits bs :: int)" + "bs \ [] \ of_bits (True # bs) = 1 + 2 * (of_bits bs :: int)" + by (simp_all add: of_bits_int_def push_bit_of_1) + +instance proof + show "of_bits (bits_of k) = k" for k :: int + by (induction k rule: int_bit_induct) simp_all +qed + +lemma bits_of_1_int [simp]: + "bits_of (1 :: int) = [True, False]" + by simp + +lemma bits_of_int_numeral_simps [simp]: + "bits_of (numeral Num.One :: int) = [True, False]" (is ?One) + "bits_of (numeral (Num.Bit0 n) :: int) = False # bits_of (numeral n :: int)" (is ?Bit0) + "bits_of (numeral (Num.Bit1 n) :: int) = True # bits_of (numeral n :: int)" (is ?Bit1) + "bits_of (- numeral (Num.Bit0 n) :: int) = False # bits_of (- numeral n :: int)" (is ?nBit0) + "bits_of (- numeral (Num.Bit1 n) :: int) = True # bits_of (- numeral (Num.inc n) :: int)" (is ?nBit1) +proof - + show ?One + by simp + define k :: int where "k = numeral n" + then have "k > 0" and *: "numeral n = k" "numeral (Num.Bit0 n) = 2 * k" "numeral (Num.Bit1 n) = 2 * k + 1" + "numeral (Num.inc n) = k + 1" + by (simp_all add: add_One) + have "- (2 * k) div 2 = - k" "(- (2 * k) - 1) div 2 = - k - 1" + by simp_all + with \k > 0\ show ?Bit0 ?Bit1 ?nBit0 ?nBit1 + by (simp_all add: *) +qed + +lemma of_bits_append [simp]: + "of_bits (bs @ cs) = of_bits bs + push_bit (length bs) (of_bits cs :: int)" + if "bs \ []" "\ last bs" +using that proof (induction bs rule: list_nonempty_induct) + case (single b) + then show ?case + by simp +next + case (cons b bs) + then show ?case + by (cases b) (simp_all add: push_bit_double) +qed + +lemma of_bits_replicate_False [simp]: + "of_bits (replicate n False) = (0 :: int)" + by (auto simp add: of_bits_int_def) + +lemma of_bits_drop [simp]: + "of_bits (drop n bs) = drop_bit n (of_bits bs :: int)" + if "n < length bs" +using that proof (induction bs arbitrary: n) + case Nil + then show ?case + by simp +next + case (Cons b bs) + show ?case + proof (cases n) + case 0 + then show ?thesis + by simp + next + case (Suc n) + with Cons.prems have "bs \ []" + by auto + with Suc Cons.IH [of n] Cons.prems show ?thesis + by (cases b) simp_all + qed +qed end -class semiring_bits = unique_euclidean_semiring_with_nat + - assumes half_measure: "a div 2 \ a \ euclidean_size (a div 2) < euclidean_size a" - \ \It is not clear whether this could be derived from already existing assumptions.\ + +subsection \Syntactic bit operations\ + +class bit_operations = bit_representation + + fixes not :: "'a \ 'a" ("NOT _" [70] 71) + and "and" :: "'a \ 'a \ 'a" (infixr "AND" 64) + and or :: "'a \ 'a \ 'a" (infixr "OR" 59) + and xor :: "'a \ 'a \ 'a" (infixr "XOR" 59) + and shift_left :: "'a \ nat \ 'a" (infixl "<<" 55) + and shift_right :: "'a \ nat \ 'a" (infixl ">>" 55) + assumes not_eq: "not = of_bits \ map Not \ bits_of" + and and_eq: "length bs = length cs \ + of_bits bs AND of_bits cs = of_bits (map2 (\) bs cs)" + and semilattice_and: "semilattice (AND)" + and or_eq: "length bs = length cs \ + of_bits bs OR of_bits cs = of_bits (map2 (\) bs cs)" + and semilattice_or: "semilattice (OR)" + and xor_eq: "length bs = length cs \ + of_bits bs XOR of_bits cs = of_bits (map2 (\) bs cs)" + and abel_semigroup_xor: "abel_semigroup (XOR)" + and shift_right_eq: "a << n = of_bits (replicate n False @ bits_of a)" + and shift_left_eq: "n < length (bits_of a) \ a >> n = of_bits (drop n (bits_of a))" begin -function bits_of :: "'a \ bool list" - where "bits_of a = odd a # (let b = a div 2 in if a = b then [] else bits_of b)" +text \ + We want the bitwise operations to bind slightly weaker + than \+\ and \-\, but \~~\ to + bind slightly stronger than \*\. +\ + +sublocale "and": semilattice "(AND)" + by (fact semilattice_and) + +sublocale or: semilattice "(OR)" + by (fact semilattice_or) + +sublocale xor: abel_semigroup "(XOR)" + by (fact abel_semigroup_xor) + +end + + +subsubsection \Instance \<^typ>\nat\\ + +locale zip_nat = single: abel_semigroup f + for f :: "bool \ bool \ bool" (infixl "\<^bold>*" 70) + + assumes end_of_bits: "\ False \<^bold>* False" +begin + +lemma False_P_imp: + "False \<^bold>* True \ P" if "False \<^bold>* P" + using that end_of_bits by (cases P) simp_all + +function F :: "nat \ nat \ nat" (infixl "\<^bold>\" 70) + where "m \<^bold>\ n = (if m = 0 \ n = 0 then 0 + else of_bool (odd m \<^bold>* odd n) + (m div 2) \<^bold>\ (n div 2) * 2)" by auto termination - by (relation "measure euclidean_size") (auto intro: half_measure) + by (relation "measure (case_prod (+))") auto -lemma bits_of_not_empty [simp]: - "bits_of a \ []" - by (induction a rule: bits_of.induct) simp_all +lemma zero_left_eq: + "0 \<^bold>\ n = of_bool (False \<^bold>* True) * n" + by (induction n rule: nat_bit_induct) (simp_all add: end_of_bits) + +lemma zero_right_eq: + "m \<^bold>\ 0 = of_bool (True \<^bold>* False) * m" + by (induction m rule: nat_bit_induct) (simp_all add: end_of_bits) -lemma bits_of_0 [simp]: - "bits_of 0 = [False]" - by simp +lemma simps [simp]: + "0 \<^bold>\ 0 = 0" + "0 \<^bold>\ n = of_bool (False \<^bold>* True) * n" + "m \<^bold>\ 0 = of_bool (True \<^bold>* False) * m" + "m > 0 \ n > 0 \ m \<^bold>\ n = of_bool (odd m \<^bold>* odd n) + (m div 2) \<^bold>\ (n div 2) * 2" + by (simp_all only: zero_left_eq zero_right_eq) simp -lemma bits_of_1 [simp]: - "bits_of 1 = [True, False]" - by simp +lemma rec: + "m \<^bold>\ n = of_bool (odd m \<^bold>* odd n) + (m div 2) \<^bold>\ (n div 2) * 2" + by (cases "m = 0 \ n = 0") (auto simp add: end_of_bits) + +declare F.simps [simp del] -lemma bits_of_double [simp]: - "bits_of (a * 2) = False # (if a = 0 then [] else bits_of a)" - by simp (simp add: mult_2_right) - -lemma bits_of_add_1_double [simp]: - "bits_of (1 + a * 2) = True # (if a + 1 = 0 then [] else bits_of a)" - by simp (simp add: mult_2_right algebra_simps) - -declare bits_of.simps [simp del] +sublocale abel_semigroup F +proof + show "m \<^bold>\ n \<^bold>\ q = m \<^bold>\ (n \<^bold>\ q)" for m n q :: nat + proof (induction m arbitrary: n q rule: nat_bit_induct) + case zero + show ?case + by simp + next + case (even m) + with rec [of "2 * m"] rec [of _ q] show ?case + by (cases "even n") (auto dest: False_P_imp) + next + case (odd m) + with rec [of "Suc (2 * m)"] rec [of _ q] show ?case + by (cases "even n"; cases "even q") + (auto dest: False_P_imp simp add: ac_simps) + qed + show "m \<^bold>\ n = n \<^bold>\ m" for m n :: nat + proof (induction m arbitrary: n rule: nat_bit_induct) + case zero + show ?case + by (simp add: ac_simps) + next + case (even m) + with rec [of "2 * m" n] rec [of n "2 * m"] show ?case + by (simp add: ac_simps) + next + case (odd m) + with rec [of "Suc (2 * m)" n] rec [of n "Suc (2 * m)"] show ?case + by (simp add: ac_simps) + qed +qed -lemma not_last_bits_of_nat [simp]: - "\ last (bits_of (of_nat n))" - by (induction n rule: nat_bit_induct) - (use of_nat_neq_0 in \simp_all add: algebra_simps\) +lemma self [simp]: + "n \<^bold>\ n = of_bool (True \<^bold>* True) * n" + by (induction n rule: nat_bit_induct) (simp_all add: end_of_bits) -lemma of_unsigned_bits_of_nat: - "of_unsigned (bits_of (of_nat n)) = of_nat n" - by (induction n rule: nat_bit_induct) - (use of_nat_neq_0 in \simp_all add: algebra_simps\) +lemma even_iff [simp]: + "even (m \<^bold>\ n) \ \ (odd m \<^bold>* odd n)" +proof (induction m arbitrary: n rule: nat_bit_induct) + case zero + show ?case + by (cases "even n") (simp_all add: end_of_bits) +next + case (even m) + then show ?case + by (simp add: rec [of "2 * m"]) +next + case (odd m) + then show ?case + by (simp add: rec [of "Suc (2 * m)"]) +qed + +lemma of_bits: + "of_bits bs \<^bold>\ of_bits cs = (of_bits (map2 (\<^bold>*) bs cs) :: nat)" + if "length bs = length cs" for bs cs +using that proof (induction bs cs rule: list_induct2) + case Nil + then show ?case + by simp +next + case (Cons b bs c cs) + then show ?case + by (cases "of_bits bs = (0::nat) \ of_bits cs = (0::nat)") + (auto simp add: ac_simps end_of_bits rec [of "Suc 0"]) +qed end -instance nat :: semiring_bits - by standard simp - -lemma bits_of_Suc_double [simp]: - "bits_of (Suc (n * 2)) = True # bits_of n" - using bits_of_add_1_double [of n] by simp - -lemma of_unsigned_bits_of: - "of_unsigned (bits_of n) = n" for n :: nat - using of_unsigned_bits_of_nat [of n, where ?'a = nat] by simp - -class ring_bits = unique_euclidean_ring_with_nat + semiring_bits +instantiation nat :: bit_operations begin -lemma bits_of_minus_1 [simp]: - "bits_of (- 1) = [True]" - using bits_of.simps [of "- 1"] by simp +definition not_nat :: "nat \ nat" + where "not_nat = of_bits \ map Not \ bits_of" + +global_interpretation and_nat: zip_nat "(\)" + defines and_nat = and_nat.F + by standard auto -lemma bits_of_double [simp]: - "bits_of (- (a * 2)) = False # (if a = 0 then [] else bits_of (- a))" - using bits_of.simps [of "- (a * 2)"] nonzero_mult_div_cancel_right [of 2 "- a"] - by simp (simp add: mult_2_right) +global_interpretation and_nat: semilattice "(AND) :: nat \ nat \ nat" +proof (rule semilattice.intro, fact and_nat.abel_semigroup_axioms, standard) + show "n AND n = n" for n :: nat + by (simp add: and_nat.self) +qed + +declare and_nat.simps [simp] \ \inconsistent declaration handling by + \global_interpretation\ in \instantiation\\ -lemma bits_of_minus_1_diff_double [simp]: - "bits_of (- 1 - a * 2) = True # (if a = 0 then [] else bits_of (- 1 - a))" -proof - - have [simp]: "- 1 - a = - a - 1" - by (simp add: algebra_simps) - show ?thesis - using bits_of.simps [of "- 1 - a * 2"] div_mult_self1 [of 2 "- 1" "- a"] - by simp (simp add: mult_2_right algebra_simps) +lemma zero_nat_and_eq [simp]: + "0 AND n = 0" for n :: nat + by simp + +lemma and_zero_nat_eq [simp]: + "n AND 0 = 0" for n :: nat + by simp + +global_interpretation or_nat: zip_nat "(\)" + defines or_nat = or_nat.F + by standard auto + +global_interpretation or_nat: semilattice "(OR) :: nat \ nat \ nat" +proof (rule semilattice.intro, fact or_nat.abel_semigroup_axioms, standard) + show "n OR n = n" for n :: nat + by (simp add: or_nat.self) qed -lemma last_bits_of_neg_of_nat [simp]: - "last (bits_of (- 1 - of_nat n))" -proof (induction n rule: nat_bit_induct) - case zero - then show ?case +declare or_nat.simps [simp] \ \inconsistent declaration handling by + \global_interpretation\ in \instantiation\\ + +lemma zero_nat_or_eq [simp]: + "0 OR n = n" for n :: nat + by simp + +lemma or_zero_nat_eq [simp]: + "n OR 0 = n" for n :: nat + by simp + +global_interpretation xor_nat: zip_nat "(\)" + defines xor_nat = xor_nat.F + by standard auto + +declare xor_nat.simps [simp] \ \inconsistent declaration handling by + \global_interpretation\ in \instantiation\\ + +lemma zero_nat_xor_eq [simp]: + "0 XOR n = n" for n :: nat + by simp + +lemma xor_zero_nat_eq [simp]: + "n XOR 0 = n" for n :: nat + by simp + +definition shift_left_nat :: "nat \ nat \ nat" + where [simp]: "m << n = push_bit n m" for m :: nat + +definition shift_right_nat :: "nat \ nat \ nat" + where [simp]: "m >> n = drop_bit n m" for m :: nat + +instance proof + show "semilattice ((AND) :: nat \ _)" + by (fact and_nat.semilattice_axioms) + show "semilattice ((OR):: nat \ _)" + by (fact or_nat.semilattice_axioms) + show "abel_semigroup ((XOR):: nat \ _)" + by (fact xor_nat.abel_semigroup_axioms) + show "(not :: nat \ _) = of_bits \ map Not \ bits_of" + by (fact not_nat_def) + show "of_bits bs AND of_bits cs = (of_bits (map2 (\) bs cs) :: nat)" + if "length bs = length cs" for bs cs + using that by (fact and_nat.of_bits) + show "of_bits bs OR of_bits cs = (of_bits (map2 (\) bs cs) :: nat)" + if "length bs = length cs" for bs cs + using that by (fact or_nat.of_bits) + show "of_bits bs XOR of_bits cs = (of_bits (map2 (\) bs cs) :: nat)" + if "length bs = length cs" for bs cs + using that by (fact xor_nat.of_bits) + show "m << n = of_bits (replicate n False @ bits_of m)" + for m n :: nat by simp -next - case (even n) - then show ?case - by (simp add: algebra_simps) -next - case (odd n) - then have "last (bits_of ((- 1 - of_nat n) * 2))" - by auto - also have "(- 1 - of_nat n) * 2 = - 1 - (1 + 2 * of_nat n)" - by (simp add: algebra_simps) - finally show ?case + show "m >> n = of_bits (drop n (bits_of m))" + for m n :: nat by simp qed -lemma of_signed_bits_of_nat: - "of_signed (bits_of (of_nat n)) = of_nat n" - by (simp add: of_signed_def of_unsigned_bits_of_nat) +end + +global_interpretation or_nat: semilattice_neutr "(OR)" "0 :: nat" + by standard simp + +global_interpretation xor_nat: comm_monoid "(XOR)" "0 :: nat" + by standard simp + +lemma not_nat_simps [simp]: + "NOT 0 = (0::nat)" + "n > 0 \ NOT n = of_bool (even n) + 2 * NOT (n div 2)" for n :: nat + by (simp_all add: not_eq) + +lemma not_1_nat [simp]: + "NOT 1 = (0::nat)" + by simp + +lemma not_Suc_0 [simp]: + "NOT (Suc 0) = (0::nat)" + by simp + +lemma Suc_0_and_eq [simp]: + "Suc 0 AND n = n mod 2" + by (cases n) auto + +lemma and_Suc_0_eq [simp]: + "n AND Suc 0 = n mod 2" + using Suc_0_and_eq [of n] by (simp add: ac_simps) + +lemma Suc_0_or_eq [simp]: + "Suc 0 OR n = n + of_bool (even n)" + by (cases n) (simp_all add: ac_simps) + +lemma or_Suc_0_eq [simp]: + "n OR Suc 0 = n + of_bool (even n)" + using Suc_0_or_eq [of n] by (simp add: ac_simps) + +lemma Suc_0_xor_eq [simp]: + "Suc 0 XOR n = n + of_bool (even n) - of_bool (odd n)" + by (cases n) (simp_all add: ac_simps) + +lemma xor_Suc_0_eq [simp]: + "n XOR Suc 0 = n + of_bool (even n) - of_bool (odd n)" + using Suc_0_xor_eq [of n] by (simp add: ac_simps) + + +subsubsection \Instance \<^typ>\int\\ + +abbreviation (input) complement :: "int \ int" + where "complement k \ - k - 1" + +lemma complement_half: + "complement (k * 2) div 2 = complement k" + by simp + +locale zip_int = single: abel_semigroup f + for f :: "bool \ bool \ bool" (infixl "\<^bold>*" 70) +begin + +lemma False_False_imp_True_True: + "True \<^bold>* True" if "False \<^bold>* False" +proof (rule ccontr) + assume "\ True \<^bold>* True" + with that show False + using single.assoc [of False True True] + by (cases "False \<^bold>* True") simp_all +qed + +function F :: "int \ int \ int" (infixl "\<^bold>\" 70) + where "k \<^bold>\ l = (if k \ {0, - 1} \ l \ {0, - 1} + then - of_bool (odd k \<^bold>* odd l) + else of_bool (odd k \<^bold>* odd l) + (k div 2) \<^bold>\ (l div 2) * 2)" + by auto + +termination + by (relation "measure (\(k, l). nat (\k\ + \l\))") auto + +lemma zero_left_eq: + "0 \<^bold>\ l = (case (False \<^bold>* False, False \<^bold>* True) + of (False, False) \ 0 + | (False, True) \ l + | (True, False) \ complement l + | (True, True) \ - 1)" + by (induction l rule: int_bit_induct) + (simp_all split: bool.split) + +lemma minus_left_eq: + "- 1 \<^bold>\ l = (case (True \<^bold>* False, True \<^bold>* True) + of (False, False) \ 0 + | (False, True) \ l + | (True, False) \ complement l + | (True, True) \ - 1)" + by (induction l rule: int_bit_induct) + (simp_all split: bool.split) + +lemma zero_right_eq: + "k \<^bold>\ 0 = (case (False \<^bold>* False, False \<^bold>* True) + of (False, False) \ 0 + | (False, True) \ k + | (True, False) \ complement k + | (True, True) \ - 1)" + by (induction k rule: int_bit_induct) + (simp_all add: ac_simps split: bool.split) -lemma of_signed_bits_neg_of_nat: - "of_signed (bits_of (- 1 - of_nat n)) = - 1 - of_nat n" -proof - - have "of_unsigned (map Not (bits_of (- 1 - of_nat n))) = of_nat n" - proof (induction n rule: nat_bit_induct) +lemma minus_right_eq: + "k \<^bold>\ - 1 = (case (True \<^bold>* False, True \<^bold>* True) + of (False, False) \ 0 + | (False, True) \ k + | (True, False) \ complement k + | (True, True) \ - 1)" + by (induction k rule: int_bit_induct) + (simp_all add: ac_simps split: bool.split) + +lemma simps [simp]: + "0 \<^bold>\ 0 = - of_bool (False \<^bold>* False)" + "- 1 \<^bold>\ 0 = - of_bool (True \<^bold>* False)" + "0 \<^bold>\ - 1 = - of_bool (False \<^bold>* True)" + "- 1 \<^bold>\ - 1 = - of_bool (True \<^bold>* True)" + "0 \<^bold>\ l = (case (False \<^bold>* False, False \<^bold>* True) + of (False, False) \ 0 + | (False, True) \ l + | (True, False) \ complement l + | (True, True) \ - 1)" + "- 1 \<^bold>\ l = (case (True \<^bold>* False, True \<^bold>* True) + of (False, False) \ 0 + | (False, True) \ l + | (True, False) \ complement l + | (True, True) \ - 1)" + "k \<^bold>\ 0 = (case (False \<^bold>* False, False \<^bold>* True) + of (False, False) \ 0 + | (False, True) \ k + | (True, False) \ complement k + | (True, True) \ - 1)" + "k \<^bold>\ - 1 = (case (True \<^bold>* False, True \<^bold>* True) + of (False, False) \ 0 + | (False, True) \ k + | (True, False) \ complement k + | (True, True) \ - 1)" + "k \ 0 \ k \ - 1 \ l \ 0 \ l \ - 1 + \ k \<^bold>\ l = of_bool (odd k \<^bold>* odd l) + (k div 2) \<^bold>\ (l div 2) * 2" + by simp_all[4] (simp_all only: zero_left_eq minus_left_eq zero_right_eq minus_right_eq, simp) + +declare F.simps [simp del] + +lemma rec: + "k \<^bold>\ l = of_bool (odd k \<^bold>* odd l) + (k div 2) \<^bold>\ (l div 2) * 2" + by (cases "k \ {0, - 1} \ l \ {0, - 1}") (auto simp add: ac_simps F.simps [of k l] split: bool.split) + +sublocale abel_semigroup F +proof + show "k \<^bold>\ l \<^bold>\ r = k \<^bold>\ (l \<^bold>\ r)" for k l r :: int + proof (induction k arbitrary: l r rule: int_bit_induct) case zero + have "complement l \<^bold>\ r = complement (l \<^bold>\ r)" if "False \<^bold>* False" "\ False \<^bold>* True" + proof (induction l arbitrary: r rule: int_bit_induct) + case zero + from that show ?case + by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits) + next + case minus + from that show ?case + by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits) + next + case (even l) + with that rec [of _ r] show ?case + by (cases "even r") + (auto simp add: complement_half ac_simps False_False_imp_True_True split: bool.splits) + next + case (odd l) + moreover have "- l - 1 = - 1 - l" + by simp + ultimately show ?case + using that rec [of _ r] + by (cases "even r") + (auto simp add: ac_simps False_False_imp_True_True split: bool.splits) + qed then show ?case + by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits) + next + case minus + have "complement l \<^bold>\ r = complement (l \<^bold>\ r)" if "\ True \<^bold>* True" "False \<^bold>* True" + proof (induction l arbitrary: r rule: int_bit_induct) + case zero + from that show ?case + by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits) + next + case minus + from that show ?case + by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits) + next + case (even l) + with that rec [of _ r] show ?case + by (cases "even r") + (auto simp add: complement_half ac_simps False_False_imp_True_True split: bool.splits) + next + case (odd l) + moreover have "- l - 1 = - 1 - l" + by simp + ultimately show ?case + using that rec [of _ r] + by (cases "even r") + (auto simp add: ac_simps False_False_imp_True_True split: bool.splits) + qed + then show ?case + by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits) + next + case (even k) + with rec [of "k * 2"] rec [of _ r] show ?case + by (cases "even r"; cases "even l") (auto simp add: ac_simps False_False_imp_True_True) + next + case (odd k) + with rec [of "1 + k * 2"] rec [of _ r] show ?case + by (cases "even r"; cases "even l") (auto simp add: ac_simps False_False_imp_True_True) + qed + show "k \<^bold>\ l = l \<^bold>\ k" for k l :: int + proof (induction k arbitrary: l rule: int_bit_induct) + case zero + show ?case by simp next - case (even n) - then show ?case - by (simp add: algebra_simps) + case minus + show ?case + by simp next - case (odd n) - have *: "- 1 - (1 + of_nat n * 2) = - 2 - of_nat n * 2" - by (simp add: algebra_simps) (metis add_assoc one_add_one) - from odd show ?case - using bits_of_double [of "of_nat (Suc n)"] of_nat_neq_0 - by (simp add: algebra_simps *) + case (even k) + with rec [of "k * 2" l] rec [of l "k * 2"] show ?case + by (simp add: ac_simps) + next + case (odd k) + with rec [of "k * 2 + 1" l] rec [of l "k * 2 + 1"] show ?case + by (simp add: ac_simps) qed - then show ?thesis - by (simp add: of_signed_def algebra_simps) qed -lemma of_signed_bits_of_int: - "of_signed (bits_of (of_int k)) = of_int k" - by (cases k rule: int_cases) - (simp_all add: of_signed_bits_of_nat of_signed_bits_neg_of_nat) +lemma self [simp]: + "k \<^bold>\ k = (case (False \<^bold>* False, True \<^bold>* True) + of (False, False) \ 0 + | (False, True) \ k + | (True, True) \ - 1)" + by (induction k rule: int_bit_induct) (auto simp add: False_False_imp_True_True split: bool.split) + +lemma even_iff [simp]: + "even (k \<^bold>\ l) \ \ (odd k \<^bold>* odd l)" +proof (induction k arbitrary: l rule: int_bit_induct) + case zero + show ?case + by (cases "even l") (simp_all split: bool.splits) +next + case minus + show ?case + by (cases "even l") (simp_all split: bool.splits) +next + case (even k) + then show ?case + by (simp add: rec [of "k * 2"]) +next + case (odd k) + then show ?case + by (simp add: rec [of "1 + k * 2"]) +qed + +lemma of_bits: + "of_bits bs \<^bold>\ of_bits cs = (of_bits (map2 (\<^bold>*) bs cs) :: int)" + if "length bs = length cs" and "\ False \<^bold>* False" for bs cs +using \length bs = length cs\ proof (induction bs cs rule: list_induct2) + case Nil + then show ?case + using \\ False \<^bold>* False\ by (auto simp add: False_False_imp_True_True split: bool.splits) +next + case (Cons b bs c cs) + show ?case + proof (cases "bs = []") + case True + with Cons show ?thesis + using \\ False \<^bold>* False\ by (cases b; cases c) + (auto simp add: ac_simps split: bool.splits) + next + case False + with Cons.hyps have "cs \ []" + by auto + with \bs \ []\ have "map2 (\<^bold>*) bs cs \ []" + by (simp add: zip_eq_Nil_iff) + from \bs \ []\ \cs \ []\ \map2 (\<^bold>*) bs cs \ []\ Cons show ?thesis + by (cases b; cases c) + (auto simp add: \\ False \<^bold>* False\ ac_simps + rec [of "of_bits bs * 2"] rec [of "of_bits cs * 2"] + rec [of "1 + of_bits bs * 2"]) + qed +qed end -instance int :: ring_bits - by standard auto +instantiation int :: bit_operations +begin + +definition not_int :: "int \ int" + where "not_int = complement" + +global_interpretation and_int: zip_int "(\)" + defines and_int = and_int.F + by standard + +declare and_int.simps [simp] \ \inconsistent declaration handling by + \global_interpretation\ in \instantiation\\ + +global_interpretation and_int: semilattice "(AND) :: int \ int \ int" +proof (rule semilattice.intro, fact and_int.abel_semigroup_axioms, standard) + show "k AND k = k" for k :: int + by (simp add: and_int.self) +qed + +lemma zero_int_and_eq [simp]: + "0 AND k = 0" for k :: int + by simp + +lemma and_zero_int_eq [simp]: + "k AND 0 = 0" for k :: int + by simp + +lemma minus_int_and_eq [simp]: + "- 1 AND k = k" for k :: int + by simp + +lemma and_minus_int_eq [simp]: + "k AND - 1 = k" for k :: int + by simp + +global_interpretation or_int: zip_int "(\)" + defines or_int = or_int.F + by standard + +declare or_int.simps [simp] \ \inconsistent declaration handling by + \global_interpretation\ in \instantiation\\ + +global_interpretation or_int: semilattice "(OR) :: int \ int \ int" +proof (rule semilattice.intro, fact or_int.abel_semigroup_axioms, standard) + show "k OR k = k" for k :: int + by (simp add: or_int.self) +qed + +lemma zero_int_or_eq [simp]: + "0 OR k = k" for k :: int + by simp + +lemma and_zero_or_eq [simp]: + "k OR 0 = k" for k :: int + by simp + +lemma minus_int_or_eq [simp]: + "- 1 OR k = - 1" for k :: int + by simp -lemma of_signed_bits_of: - "of_signed (bits_of k) = k" for k :: int - using of_signed_bits_of_int [of k, where ?'a = int] by simp +lemma or_minus_int_eq [simp]: + "k OR - 1 = - 1" for k :: int + by simp + +global_interpretation xor_int: zip_int "(\)" + defines xor_int = xor_int.F + by standard + +declare xor_int.simps [simp] \ \inconsistent declaration handling by + \global_interpretation\ in \instantiation\\ + +lemma zero_int_xor_eq [simp]: + "0 XOR k = k" for k :: int + by simp + +lemma and_zero_xor_eq [simp]: + "k XOR 0 = k" for k :: int + by simp + +lemma minus_int_xor_eq [simp]: + "- 1 XOR k = complement k" for k :: int + by simp + +lemma xor_minus_int_eq [simp]: + "k XOR - 1 = complement k" for k :: int + by simp + +definition shift_left_int :: "int \ nat \ int" + where [simp]: "k << n = push_bit n k" for k :: int + +definition shift_right_int :: "int \ nat \ int" + where [simp]: "k >> n = drop_bit n k" for k :: int + +instance proof + show "semilattice ((AND) :: int \ _)" + by (fact and_int.semilattice_axioms) + show "semilattice ((OR) :: int \ _)" + by (fact or_int.semilattice_axioms) + show "abel_semigroup ((XOR) :: int \ _)" + by (fact xor_int.abel_semigroup_axioms) + show "(not :: int \ _) = of_bits \ map Not \ bits_of" + proof (rule sym, rule ext) + fix k :: int + show "(of_bits \ map Not \ bits_of) k = NOT k" + by (induction k rule: int_bit_induct) (simp_all add: not_int_def) + qed + show "of_bits bs AND of_bits cs = (of_bits (map2 (\) bs cs) :: int)" + if "length bs = length cs" for bs cs + using that by (rule and_int.of_bits) simp + show "of_bits bs OR of_bits cs = (of_bits (map2 (\) bs cs) :: int)" + if "length bs = length cs" for bs cs + using that by (rule or_int.of_bits) simp + show "of_bits bs XOR of_bits cs = (of_bits (map2 (\) bs cs) :: int)" + if "length bs = length cs" for bs cs + using that by (rule xor_int.of_bits) simp + show "k << n = of_bits (replicate n False @ bits_of k)" + for k :: int and n :: nat + by (cases "n = 0") simp_all + show "k >> n = of_bits (drop n (bits_of k))" + if "n < length (bits_of k)" + for k :: int and n :: nat + using that by simp +qed end + +global_interpretation and_int: semilattice_neutr "(AND)" "- 1 :: int" + by standard simp + +global_interpretation or_int: semilattice_neutr "(OR)" "0 :: int" + by standard simp + +global_interpretation xor_int: comm_monoid "(XOR)" "0 :: int" + by standard simp + +lemma not_int_simps [simp]: + "NOT 0 = (- 1 :: int)" + "NOT - 1 = (0 :: int)" + "k \ 0 \ k \ - 1 \ NOT k = of_bool (even k) + 2 * NOT (k div 2)" for k :: int + by (auto simp add: not_int_def elim: oddE) + +lemma not_one_int [simp]: + "NOT 1 = (- 2 :: int)" + by simp + +lemma one_and_int_eq [simp]: + "1 AND k = k mod 2" for k :: int + using and_int.rec [of 1] by (simp add: mod2_eq_if) + +lemma and_one_int_eq [simp]: + "k AND 1 = k mod 2" for k :: int + using one_and_int_eq [of 1] by (simp add: ac_simps) + +lemma one_or_int_eq [simp]: + "1 OR k = k + of_bool (even k)" for k :: int + using or_int.rec [of 1] by (auto elim: oddE) + +lemma or_one_int_eq [simp]: + "k OR 1 = k + of_bool (even k)" for k :: int + using one_or_int_eq [of k] by (simp add: ac_simps) + +lemma one_xor_int_eq [simp]: + "1 XOR k = k + of_bool (even k) - of_bool (odd k)" for k :: int + using xor_int.rec [of 1] by (auto elim: oddE) + +lemma xor_one_int_eq [simp]: + "k XOR 1 = k + of_bool (even k) - of_bool (odd k)" for k :: int + using one_xor_int_eq [of k] by (simp add: ac_simps) + +end