Skip to content

Commit

Permalink
feat: add BitVec comparison lemmas to bv_normalize (#6799)
Browse files Browse the repository at this point in the history
This PR adds a number of simple comparison lemmas to the top/bottom
element for BitVec. Then they are applied to teach bv_normalize that
`(a<1) = (a==0)` and to remove an intermediate proof that is no longer
necessary along the way.
  • Loading branch information
vlad902 authored and luisacicolini committed Jan 28, 2025
1 parent 6595baf commit c0f8403
Show file tree
Hide file tree
Showing 3 changed files with 43 additions and 17 deletions.
39 changes: 33 additions & 6 deletions src/Init/Data/BitVec/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -2716,12 +2716,39 @@ theorem not_lt_iff_le {x y : BitVec w} : (¬ x < y) ↔ y ≤ x := by
constructor <;>
(intro h; simp only [lt_def, Nat.not_lt, le_def] at h ⊢; omega)

theorem zero_lt_of_neq_zero {x : BitVec w} (hx : x ≠ 0#w) :
0#w < x := by
simp only [lt_def, toNat_ofNat, Nat.zero_mod]
have := (BitVec.toNat_eq (x := x) (y := 0#w)).subst hx
simp [toNat_ofNat, Nat.zero_mod, gt_iff_lt] at this ⊢
omega
@[simp]
theorem not_lt_zero {x : BitVec w} : ¬x < 0#w := of_decide_eq_false rfl

@[simp]
theorem le_zero_iff {x : BitVec w} : x ≤ 0#w ↔ x = 0#w := by
constructor
· intro h
have : x ≥ 0 := not_lt_iff_le.mp not_lt_zero
exact Eq.symm (BitVec.le_antisymm this h)
· simp_all

@[simp]
theorem lt_one_iff {x : BitVec w} (h : 0 < w) : x < 1#w ↔ x = 0#w := by
constructor
· intro h₂
rw [lt_def, toNat_ofNat, ← Int.ofNat_lt, Int.ofNat_emod, Int.ofNat_one, Int.natCast_pow,
Int.ofNat_two, @Int.emod_eq_of_lt 1 (2^w) (by omega) (by omega)] at h₂
simp [toNat_eq, show x.toNat = 0 by omega]
· simp_all

@[simp]
theorem not_allOnes_lt {x : BitVec w} : ¬allOnes w < x := by
have : 2^w ≠ 0 := Ne.symm (NeZero.ne' (2^w))
rw [BitVec.not_lt, le_def, Nat.le_iff_lt_add_one, toNat_allOnes, Nat.sub_one_add_one this]
exact isLt x

@[simp]
theorem allOnes_le_iff {x : BitVec w} : allOnes w ≤ x ↔ x = allOnes w := by
constructor
· intro h
have : x ≤ allOnes w := not_lt_iff_le.mp not_allOnes_lt
exact Eq.symm (BitVec.le_antisymm h this)
· simp_all

/-! ### udiv -/

Expand Down
17 changes: 6 additions & 11 deletions src/Std/Tactic/BVDecide/Normalize/BitVec.lean
Original file line number Diff line number Diff line change
Expand Up @@ -299,20 +299,15 @@ theorem BitVec.lt_irrefl (a : BitVec n) : (BitVec.ult a a) = false := by
@[bv_normalize]
theorem BitVec.not_lt_zero (a : BitVec n) : (BitVec.ult a 0#n) = false := by rfl

theorem BitVec.max_ult (a : BitVec w) : ¬ ((-1#w) < a) := by
rcases w with rfl | w
· simp [bv_toNat, BitVec.toNat_of_zero_length]
· simp only [BitVec.lt_def, BitVec.toNat_neg, BitVec.toNat_ofNat, Nat.not_lt]
rw [Nat.mod_eq_of_lt (a := 1) (by simp)];
rw [Nat.mod_eq_of_lt]
· omega
· apply Nat.sub_one_lt_of_le (Nat.pow_pos (by omega)) (Nat.le_refl ..)
@[bv_normalize]
theorem BitVec.lt_one_iff (a : BitVec n) (h : 0 < n) : (BitVec.ult a 1#n) = (a == 0#n) := by
rw [Bool.eq_iff_iff, beq_iff_eq, ← BitVec.lt_ult]
exact _root_.BitVec.lt_one_iff h

-- used in simproc because of -1#w normalisation
theorem BitVec.max_ult' (a : BitVec w) : (BitVec.ult (-1#w) a) = false := by
have := BitVec.max_ult a
rw [BitVec.lt_ult] at this
simp [this]
rw [BitVec.negOne_eq_allOnes, ← Bool.not_eq_true, ← @lt_ult]
exact BitVec.not_allOnes_lt

attribute [bv_normalize] BitVec.replicate_zero_eq
attribute [bv_normalize] BitVec.add_eq_xor
Expand Down
4 changes: 4 additions & 0 deletions tests/lean/run/bv_decide_rewriter.lean
Original file line number Diff line number Diff line change
Expand Up @@ -97,6 +97,10 @@ example (x : BitVec 16) : ¬x < 0 := by bv_normalize
example (x : BitVec 16) : x ≥ 0 := by bv_normalize
example (x : BitVec 16) : !(x.ult 0) := by bv_normalize

-- lt_one_iff
example (x : BitVec 16) : (x < 1) ↔ (x = 0) := by bv_normalize
example (x : BitVec 16) : (x.ult 1) = (x == 0) := by bv_normalize

section

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

0 comments on commit c0f8403

Please sign in to comment.