Skip to content

Commit

Permalink
feat: teach bv_normalize that (x >> x) = 0 (#6818)
Browse files Browse the repository at this point in the history
This PR adds a BitVec lemma that `(x >> x) = 0` and plumbs it through to
bv_normalize. I also move some theorems I found useful to the top of the
ushiftRight section.
  • Loading branch information
vlad902 authored Jan 28, 2025
1 parent 9f5a9a0 commit 1d94397
Show file tree
Hide file tree
Showing 3 changed files with 14 additions and 6 deletions.
13 changes: 7 additions & 6 deletions src/Init/Data/BitVec/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1391,6 +1391,11 @@ theorem shiftLeft_ofNat_eq {x : BitVec w} {k : Nat} : x <<< (BitVec.ofNat w k) =

/-! ### ushiftRight -/

@[simp] theorem ushiftRight_eq' (x : BitVec w₁) (y : BitVec w₂) :
x >>> y = x >>> y.toNat := by rfl

theorem ushiftRight_ofNat_eq {x : BitVec w} {k : Nat} : x >>> (BitVec.ofNat w k) = x >>> (k % 2^w) := rfl

@[simp, bv_toNat] theorem toNat_ushiftRight (x : BitVec n) (i : Nat) :
(x >>> i).toNat = x.toNat >>> i := rfl

Expand Down Expand Up @@ -1514,13 +1519,9 @@ theorem msb_ushiftRight {x : BitVec w} {n : Nat} :
case succ nn ih =>
simp [BitVec.ushiftRight_eq, getMsbD_ushiftRight, BitVec.msb, ih, show nn + 1 > 0 by omega]

/-! ### ushiftRight reductions from BitVec to Nat -/

@[simp]
theorem ushiftRight_eq' (x : BitVec w₁) (y : BitVec w₂) :
x >>> y = x >>> y.toNat := by rfl

theorem ushiftRight_ofNat_eq {x : BitVec w} {k : Nat} : x >>> (BitVec.ofNat w k) = x >>> (k % 2^w) := rfl
theorem ushiftRight_self (n : BitVec w) : n >>> n.toNat = 0#w := by
simp [BitVec.toNat_eq, Nat.shiftRight_eq_div_pow, Nat.lt_two_pow_self, Nat.div_eq_of_lt]

/-! ### sshiftRight -/

Expand Down
4 changes: 4 additions & 0 deletions src/Std/Tactic/BVDecide/Normalize/BitVec.lean
Original file line number Diff line number Diff line change
Expand Up @@ -278,6 +278,10 @@ theorem BitVec.ushiftRight_zero' (n : BitVec w) : n >>> 0#w' = n := by
simp only [(· >>> ·)]
simp

@[bv_normalize]
theorem BitVec.ushiftRight_self (n : BitVec w) : n >>> n = 0#w := by
simp

theorem BitVec.zero_lt_iff_zero_neq (a : BitVec w) : (0#w < a) ↔ (a ≠ 0#w) := by
constructor <;>
simp_all only [BitVec.lt_def, BitVec.toNat_ofNat, Nat.zero_mod, ne_eq, BitVec.toNat_eq] <;>
Expand Down
3 changes: 3 additions & 0 deletions tests/lean/run/bv_decide_rewriter.lean
Original file line number Diff line number Diff line change
Expand Up @@ -101,6 +101,9 @@ example (x : BitVec 16) : !(x.ult 0) := by bv_normalize
example (x : BitVec 16) : (x < 1) ↔ (x = 0) := by bv_normalize
example (x : BitVec 16) : (x.ult 1) = (x == 0) := by bv_normalize

-- ushiftRight_self
example (x : BitVec 16) : (x >>> x) == 0 := by bv_normalize

section

example (x y : BitVec 256) : x * y = y * x := by
Expand Down

0 comments on commit 1d94397

Please sign in to comment.