Skip to content

Commit

Permalink
feat: add Bitvec.[add, sub, mul]_eq_xor and width_one_cases (#5554)
Browse files Browse the repository at this point in the history
Co-authored-by: Tobias Grosser <[email protected]>
  • Loading branch information
luisacicolini and tobiasgrosser authored Oct 1, 2024
1 parent ddec533 commit 3e2bca7
Showing 1 changed file with 25 additions and 0 deletions.
25 changes: 25 additions & 0 deletions src/Init/Data/BitVec/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -463,6 +463,16 @@ theorem toInt_pos_iff {w : Nat} {x : BitVec w} :
0 ≤ BitVec.toInt x ↔ 2 * x.toNat < 2 ^ w := by
simp [toInt_eq_toNat_cond]; omega

theorem eq_zero_or_eq_one (a : BitVec 1) : a = 0#1 ∨ a = 1#1 := by
obtain ⟨a, ha⟩ := a
simp only [Nat.reducePow]
have acases : a = 0 ∨ a = 1 := by omega
rcases acases with ⟨rfl | rfl⟩
· simp
· case inr h =>
subst h
simp

/-! ### setWidth, zeroExtend and truncate -/

@[simp]
Expand Down Expand Up @@ -1910,6 +1920,11 @@ theorem shiftLeft_add_distrib {x y : BitVec w} {n : Nat} :
case succ n ih =>
simp [ih, toNat_eq, Nat.shiftLeft_eq, ← Nat.add_mul]

theorem add_eq_xor {a b : BitVec 1} : a + b = a ^^^ b := by
have ha : a = 0 ∨ a = 1 := eq_zero_or_eq_one _
have hb : b = 0 ∨ b = 1 := eq_zero_or_eq_one _
rcases ha with h | h <;> (rcases hb with h' | h' <;> (simp [h, h']))

/-! ### sub/neg -/

theorem sub_def {n} (x y : BitVec n) : x - y = .ofNat n ((2^n - y.toNat) + x.toNat) := by rfl
Expand Down Expand Up @@ -2019,6 +2034,11 @@ theorem neg_ne_iff_ne_neg {x y : BitVec w} : -x ≠ y ↔ x ≠ -y := by
subst h'
simp at h

theorem sub_eq_xor {a b : BitVec 1} : a - b = a ^^^ b := by
have ha : a = 0 ∨ a = 1 := eq_zero_or_eq_one _
have hb : b = 0 ∨ b = 1 := eq_zero_or_eq_one _
rcases ha with h | h <;> (rcases hb with h' | h' <;> (simp [h, h']))

/-! ### abs -/

@[simp, bv_toNat]
Expand Down Expand Up @@ -2086,6 +2106,11 @@ theorem ofInt_mul {n} (x y : Int) : BitVec.ofInt n (x * y) =
apply eq_of_toInt_eq
simp

theorem mul_eq_and {a b : BitVec 1} : a * b = a &&& b := by
have ha : a = 0 ∨ a = 1 := eq_zero_or_eq_one _
have hb : b = 0 ∨ b = 1 := eq_zero_or_eq_one _
rcases ha with h | h <;> (rcases hb with h' | h' <;> (simp [h, h']))

/-! ### le and lt -/

@[bv_toNat] theorem le_def {x y : BitVec n} :
Expand Down

0 comments on commit 3e2bca7

Please sign in to comment.