src/HOL/Parity.thy
changeset 67816 2249b27ab1dd
parent 67371 2d9cf74943e1
child 67828 655d03493d0f
--- a/src/HOL/Parity.thy	Sat Mar 10 20:24:00 2018 +0100
+++ b/src/HOL/Parity.thy	Sat Mar 10 19:36:59 2018 +0000
@@ -188,6 +188,10 @@
   "a mod 2 = of_bool (odd a)"
   by (auto elim: oddE)
 
+lemma of_bool_odd_eq_mod_2:
+  "of_bool (odd a) = a mod 2"
+  by (simp add: mod_2_eq_odd)
+
 lemma one_mod_2_pow_eq [simp]:
   "1 mod (2 ^ n) = of_bool (n > 0)"
 proof -
@@ -198,6 +202,10 @@
   finally show ?thesis .
 qed
 
+lemma one_div_2_pow_eq [simp]:
+  "1 div (2 ^ n) = of_bool (n = 0)"
+  using div_mult_mod_eq [of 1 "2 ^ n"] by auto
+
 lemma even_of_nat [simp]:
   "even (of_nat a) \<longleftrightarrow> even a"
 proof -
@@ -356,7 +364,7 @@
 
 subclass comm_ring_1 ..
 
-lemma even_minus [simp]:
+lemma even_minus:
   "even (- a) \<longleftrightarrow> even a"
   by (fact dvd_minus_iff)
 
@@ -589,19 +597,173 @@
 instance int :: ring_parity
   by standard (simp_all add: dvd_eq_mod_eq_0 divide_int_def division_segment_int_def)
 
-lemma even_diff_iff [simp]:
+lemma even_diff_iff:
   "even (k - l) \<longleftrightarrow> even (k + l)" for k l :: int
-  using dvd_add_times_triv_right_iff [of 2 "k - l" l] by (simp add: mult_2_right)
+  by (fact even_diff)
 
-lemma even_abs_add_iff [simp]:
+lemma even_abs_add_iff:
   "even (\<bar>k\<bar> + l) \<longleftrightarrow> even (k + l)" for k l :: int
-  by (cases "k \<ge> 0") (simp_all add: ac_simps)
+  by simp
 
-lemma even_add_abs_iff [simp]:
+lemma even_add_abs_iff:
   "even (k + \<bar>l\<bar>) \<longleftrightarrow> even (k + l)" for k l :: int
-  using even_abs_add_iff [of l k] by (simp add: ac_simps)
+  by simp
 
 lemma even_nat_iff: "0 \<le> k \<Longrightarrow> even (nat k) \<longleftrightarrow> even k"
   by (simp add: even_of_nat [of "nat k", where ?'a = int, symmetric])
 
+
+class semiring_bits = semiring_parity +
+  assumes div_mult2_eq': "a div (of_nat m * of_nat n) = a div of_nat m div of_nat n"
+    and mod_mult2_eq': "a mod (of_nat m * of_nat n) = of_nat m * (a div of_nat m mod of_nat n) + a mod of_nat m"
+  \<comment> \<open>maybe this specifications can be simplified,
+   or even derived (partially) in @{class unique_euclidean_semiring}\<close>
+begin
+
+text \<open>The primary purpose of the following operations is
+  to avoid ad-hoc simplification of concrete expressions @{term "2 ^ n"}\<close>
+
+definition bit_push :: "nat \<Rightarrow> 'a \<Rightarrow> 'a"
+  where bit_push_eq_mult: "bit_push n a = a * 2 ^ n"
+ 
+definition bit_take :: "nat \<Rightarrow> 'a \<Rightarrow> 'a"
+  where bit_take_eq_mod: "bit_take n a = a mod of_nat (2 ^ n)"
+
+definition bit_drop :: "nat \<Rightarrow> 'a \<Rightarrow> 'a"
+  where bit_drop_eq_div: "bit_drop n a = a div of_nat (2 ^ n)"
+
+lemma bit_ident:
+  "bit_push n (bit_drop n a) + bit_take n a = a"
+  using div_mult_mod_eq by (simp add: bit_push_eq_mult bit_take_eq_mod bit_drop_eq_div)
+
+lemma bit_take_bit_take [simp]:
+  "bit_take n (bit_take n a) = bit_take n a"
+  by (simp add: bit_take_eq_mod)
+  
+lemma bit_take_0 [simp]:
+  "bit_take 0 a = 0"
+  by (simp add: bit_take_eq_mod)
+
+lemma bit_take_Suc [simp]:
+  "bit_take (Suc n) a = bit_take n (a div 2) * 2 + of_bool (odd a)"
+proof -
+  have "1 + 2 * (a div 2) mod (2 * 2 ^ n) = (a div 2 * 2 + a mod 2) mod (2 * 2 ^ n)"
+    if "odd a"
+    using that mod_mult2_eq' [of "1 + 2 * (a div 2)" 2 "2 ^ n"]
+    by (simp add: ac_simps odd_iff_mod_2_eq_one mult_mod_right)
+  also have "\<dots> = a mod (2 * 2 ^ n)"
+    by (simp only: div_mult_mod_eq)
+  finally show ?thesis
+    by (simp add: bit_take_eq_mod algebra_simps mult_mod_right)
+qed
+
+lemma bit_take_of_0 [simp]:
+  "bit_take n 0 = 0"
+  by (simp add: bit_take_eq_mod)
+
+lemma bit_take_plus:
+  "bit_take n (bit_take n a + bit_take n b) = bit_take n (a + b)"
+  by (simp add: bit_take_eq_mod mod_simps)
+
+lemma bit_take_of_1_eq_0_iff [simp]:
+  "bit_take n 1 = 0 \<longleftrightarrow> n = 0"
+  by (simp add: bit_take_eq_mod)
+
+lemma bit_push_eq_0_iff [simp]:
+  "bit_push n a = 0 \<longleftrightarrow> a = 0"
+  by (simp add: bit_push_eq_mult)
+
+lemma bit_drop_0 [simp]:
+  "bit_drop 0 = id"
+  by (simp add: fun_eq_iff bit_drop_eq_div)
+
+lemma bit_drop_of_0 [simp]:
+  "bit_drop n 0 = 0"
+  by (simp add: bit_drop_eq_div)
+
+lemma bit_drop_Suc [simp]:
+  "bit_drop (Suc n) a = bit_drop n (a div 2)"
+proof (cases "even a")
+  case True
+  then obtain b where "a = 2 * b" ..
+  moreover have "bit_drop (Suc n) (2 * b) = bit_drop n b"
+    by (simp add: bit_drop_eq_div)
+  ultimately show ?thesis
+    by simp
+next
+  case False
+  then obtain b where "a = 2 * b + 1" ..
+  moreover have "bit_drop (Suc n) (2 * b + 1) = bit_drop n b"
+    using div_mult2_eq' [of "1 + b * 2" 2 "2 ^ n"]
+    by (auto simp add: bit_drop_eq_div ac_simps)
+  ultimately show ?thesis
+    by simp
+qed
+
+lemma bit_drop_half:
+  "bit_drop n (a div 2) = bit_drop n a div 2"
+  by (induction n arbitrary: a) simp_all
+
+lemma bit_drop_of_bool [simp]:
+  "bit_drop n (of_bool d) = of_bool (n = 0 \<and> d)"
+  by (cases n) simp_all
+
+lemma even_bit_take_eq [simp]:
+  "even (bit_take n a) \<longleftrightarrow> n = 0 \<or> even a"
+  by (cases n) (simp_all add: bit_take_eq_mod dvd_mod_iff)
+
+lemma bit_push_0_id [simp]:
+  "bit_push 0 = id"
+  by (simp add: fun_eq_iff bit_push_eq_mult)
+
+lemma bit_push_of_0 [simp]:
+  "bit_push n 0 = 0"
+  by (simp add: bit_push_eq_mult)
+
+lemma bit_push_of_1:
+  "bit_push n 1 = 2 ^ n"
+  by (simp add: bit_push_eq_mult)
+
+lemma bit_push_Suc [simp]:
+  "bit_push (Suc n) a = bit_push n (a * 2)"
+  by (simp add: bit_push_eq_mult ac_simps)
+
+lemma bit_push_double:
+  "bit_push n (a * 2) = bit_push n a * 2"
+  by (simp add: bit_push_eq_mult ac_simps)
+
+lemma bit_drop_bit_take [simp]:
+  "bit_drop n (bit_take n a) = 0"
+  by (simp add: bit_drop_eq_div bit_take_eq_mod)
+
+lemma bit_take_bit_drop_commute:
+  "bit_drop m (bit_take n a) = bit_take (n - m) (bit_drop m a)"
+  for m n :: nat
+proof (cases "n \<ge> m")
+  case True
+  moreover define q where "q = n - m"
+  ultimately have "n - m = q" and "n = m + q"
+    by simp_all
+  moreover have "bit_drop m (bit_take (m + q) a) = bit_take q (bit_drop m a)"
+    using mod_mult2_eq' [of a "2 ^ m" "2 ^ q"]
+    by (simp add: bit_drop_eq_div bit_take_eq_mod power_add)
+  ultimately show ?thesis
+    by simp
+next
+  case False
+  moreover define q where "q = m - n"
+  ultimately have "m - n = q" and "m = n + q"
+    by simp_all
+  moreover have "bit_drop (n + q) (bit_take n a) = 0"
+    using div_mult2_eq' [of "a mod 2 ^ n" "2 ^ n" "2 ^ q"]
+    by (simp add: bit_drop_eq_div bit_take_eq_mod power_add div_mult2_eq)
+  ultimately show ?thesis
+    by simp
+qed
+
 end
+
+instance nat :: semiring_bits
+  by standard (simp_all add: mod_Suc Suc_double_not_eq_double div_mult2_eq mod_mult2_eq)
+
+end