Skip to content

Commit 906aa1b

Browse files
feat: add Nat.[shiftLeft_or_distrib, shiftLeft_xor_distrib, shiftLeft_and_distrib, testBit_mul_two_pow, bitwise_mul_two_pow, shiftLeft_bitwise_distrib] (#6630)
This PR adds theorems `Nat.[shiftLeft_or_distrib`, shiftLeft_xor_distrib`, shiftLeft_and_distrib`, `testBit_mul_two_pow`, `bitwise_mul_two_pow`, `shiftLeft_bitwise_distrib]`, to prove `Nat.shiftLeft_or_distrib` by emulating the proof strategy of `shiftRight_and_distrib`. In particular, `Nat.shiftLeft_or_distrib` is necessary to simplify the proofs in #6476. --------- Co-authored-by: Alex Keizer <[email protected]>
1 parent f015271 commit 906aa1b

File tree

1 file changed

+26
-0
lines changed

1 file changed

+26
-0
lines changed

src/Init/Data/Nat/Bitwise/Lemmas.lean

+26
Original file line numberDiff line numberDiff line change
@@ -711,6 +711,32 @@ theorem mul_add_lt_is_or {b : Nat} (b_lt : b < 2^i) (a : Nat) : 2^i * a + b = 2^
711711
rw [mod_two_eq_one_iff_testBit_zero, testBit_shiftLeft]
712712
simp
713713

714+
theorem testBit_mul_two_pow (x i n : Nat) :
715+
(x * 2 ^ n).testBit i = (decide (n ≤ i) && x.testBit (i - n)) := by
716+
rw [← testBit_shiftLeft, shiftLeft_eq]
717+
718+
theorem bitwise_mul_two_pow (of_false_false : f false false = false := by rfl) :
719+
(bitwise f x y) * 2 ^ n = bitwise f (x * 2 ^ n) (y * 2 ^ n) := by
720+
apply Nat.eq_of_testBit_eq
721+
simp only [testBit_mul_two_pow, testBit_bitwise of_false_false, Bool.if_false_right]
722+
intro i
723+
by_cases hn : n ≤ i
724+
· simp [hn]
725+
· simp [hn, of_false_false]
726+
727+
theorem shiftLeft_bitwise_distrib {a b : Nat} (of_false_false : f false false = false := by rfl) :
728+
(bitwise f a b) <<< i = bitwise f (a <<< i) (b <<< i) := by
729+
simp [shiftLeft_eq, bitwise_mul_two_pow of_false_false]
730+
731+
theorem shiftLeft_and_distrib {a b : Nat} : (a &&& b) <<< i = a <<< i &&& b <<< i :=
732+
shiftLeft_bitwise_distrib
733+
734+
theorem shiftLeft_or_distrib {a b : Nat} : (a ||| b) <<< i = a <<< i ||| b <<< i :=
735+
shiftLeft_bitwise_distrib
736+
737+
theorem shiftLeft_xor_distrib {a b : Nat} : (a ^^^ b) <<< i = a <<< i ^^^ b <<< i :=
738+
shiftLeft_bitwise_distrib
739+
714740
@[simp] theorem decide_shiftRight_mod_two_eq_one :
715741
decide (x >>> i % 2 = 1) = x.testBit i := by
716742
simp only [testBit, one_and_eq_mod_two, mod_two_bne_zero]

0 commit comments

Comments
 (0)