Skip to content

Commit

Permalink
feat: add BitVec.toNat_[udiv|umod] and [udiv|umod]_eq
Browse files Browse the repository at this point in the history
Co-authored-by: Siddharth <siddu.druid@gmail.com>
  • Loading branch information
tobiasgrosser and bollu committed Aug 10, 2024
1 parent 95bf679 commit 2a3cd44
Showing 1 changed file with 27 additions and 0 deletions.
27 changes: 27 additions & 0 deletions src/Init/Data/BitVec/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -857,6 +857,33 @@ theorem sshiftRight_add {x : BitVec w} {m n : Nat} :
@[simp]
theorem sshiftRight_eq' (x : BitVec w) : x.sshiftRight' y = x.sshiftRight y.toNat := rfl

/-! ### udiv -/

theorem udiv_eq {x y : BitVec n} : x.udiv y = BitVec.ofNat n (x.toNat / y.toNat) := by
have h : x.toNat / y.toNat < 2 ^ n := by
exact Nat.lt_of_le_of_lt (Nat.div_le_self ..) (by omega)
simp [udiv, bv_toNat, h, Nat.mod_eq_of_lt]

@[simp, bv_toNat]
theorem toNat_udiv {x y : BitVec n} : (x.udiv y).toNat = x.toNat / y.toNat := by
simp only [udiv_eq]
by_cases h : y = 0
· simp [h]
· rw [toNat_ofNat, Nat.mod_eq_of_lt]
exact Nat.lt_of_le_of_lt (Nat.div_le_self ..) (by omega)

/-! ### umod -/

theorem umod_eq {x y : BitVec n} :
x.umod y = BitVec.ofNat n (x.toNat % y.toNat) := by
have h : x.toNat % y.toNat < 2 ^ n := by
apply Nat.lt_of_le_of_lt (Nat.mod_le _ _) x.isLt
simp [umod, bv_toNat, Nat.mod_eq_of_lt h]

@[simp, bv_toNat]
theorem toNat_umod {x y : BitVec n} :
(x.umod y).toNat = x.toNat % y.toNat := by rfl

/-! ### signExtend -/

/-- Equation theorem for `Int.sub` when both arguments are `Int.ofNat` -/
Expand Down

0 comments on commit 2a3cd44

Please sign in to comment.