Skip to content

Commit

Permalink
feat: add twoPow multiplication lemmas (#6742)
Browse files Browse the repository at this point in the history
This PR adds the lemmas that show what happens when multiplying by
`twoPow` to an arbitrary term, as well to another `twoPow`.

This will be followed up by a PR that uses these to build a simproc to
canonicalize `twoPow w i * x` and `x * twoPow w i`.
  • Loading branch information
bollu authored Jan 22, 2025
1 parent 6595ca8 commit 6befda8
Showing 1 changed file with 11 additions and 0 deletions.
11 changes: 11 additions & 0 deletions src/Init/Data/BitVec/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3328,6 +3328,11 @@ theorem mul_twoPow_eq_shiftLeft (x : BitVec w) (i : Nat) :
apply Nat.pow_dvd_pow 2 (by omega)
simp [Nat.mul_mod, hpow]

theorem twoPow_mul_eq_shiftLeft (x : BitVec w) (i : Nat) :
(twoPow w i) * x = x <<< i := by
rw [BitVec.mul_comm, mul_twoPow_eq_shiftLeft]


theorem twoPow_zero {w : Nat} : twoPow w 0 = 1#w := by
apply eq_of_toNat_eq
simp
Expand All @@ -3337,6 +3342,12 @@ theorem shiftLeft_eq_mul_twoPow (x : BitVec w) (n : Nat) :
ext i
simp [getLsbD_shiftLeft, Fin.is_lt, decide_true, Bool.true_and, mul_twoPow_eq_shiftLeft]

/-- 2^i * 2^j = 2^(i + j) with bitvectors as well -/
theorem twoPow_mul_twoPow_eq {w : Nat} (i j : Nat) : twoPow w i * twoPow w j = twoPow w (i + j) := by
apply BitVec.eq_of_toNat_eq
simp only [toNat_mul, toNat_twoPow]
rw [← Nat.mul_mod, Nat.pow_add]

/--
The unsigned division of `x` by `2^k` equals shifting `x` right by `k`,
when `k` is less than the bitwidth `w`.
Expand Down

0 comments on commit 6befda8

Please sign in to comment.