Skip to content

Commit

Permalink
feat: BitVec.toNat BitVec.signExtend (#6155)
Browse files Browse the repository at this point in the history
This PR adds `toNat` theorems for `BitVec.signExtend.`

Sign extending to a larger bitwidth depends on the msb. If the msb is
false, then the result equals the original value. If the msb is true,
then we add a value of `(2^v - 2^w)`, which arises from the sign
extension.

```lean
theorem toNat_signExtend (x : BitVec w) {v : Nat} :
    (x.signExtend v).toNat = (x.setWidth v).toNat + if x.msb then 2^v - 2^w else 0
```

Co-authored-by: Harun Khan <[email protected]>
  • Loading branch information
bollu and mhk119 authored Nov 27, 2024
1 parent 597ef8c commit 7692343
Showing 1 changed file with 6 additions and 3 deletions.
9 changes: 6 additions & 3 deletions src/Init/Data/BitVec/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1622,22 +1622,25 @@ theorem signExtend_eq (x : BitVec w) : x.signExtend w = x := by
/-- Sign extending to a larger bitwidth depends on the msb.
If the msb is false, then the result equals the original value.
If the msb is true, then we add a value of `(2^v - 2^w)`, which arises from the sign extension. -/
theorem toNat_signExtend_of_le (x : BitVec w) {v : Nat} (hv : w ≤ v) :
private theorem toNat_signExtend_of_le (x : BitVec w) {v : Nat} (hv : w ≤ v) :
(x.signExtend v).toNat = x.toNat + if x.msb then 2^v - 2^w else 0 := by
apply Nat.eq_of_testBit_eq
intro i
have ⟨k, hk⟩ := Nat.exists_eq_add_of_le hv
rw [hk, testBit_toNat, getLsbD_signExtend, Nat.pow_add, ← Nat.mul_sub_one, Nat.add_comm (x.toNat)]
by_cases hx : x.msb
· simp [hx, Nat.testBit_mul_pow_two_add _ x.isLt, testBit_toNat]
· simp only [hx, Bool.if_true_right, ↓reduceIte,
Nat.testBit_mul_pow_two_add _ x.isLt,
testBit_toNat, Nat.testBit_two_pow_sub_one]
-- Case analysis on i being in the intervals [0..w), [w..w + k), [w+k..∞)
have hi : i < w ∨ (w ≤ i ∧ i < w + k) ∨ w + k ≤ i := by omega
rcases hi with hi | hi | hi
· simp [hi]; omega
· simp [hi]; omega
· simp [hi, show ¬ (i < w + k) by omega, show ¬ (i < w) by omega]
omega
· simp [hx, Nat.testBit_mul_pow_two_add _ x.isLt, testBit_toNat]
· simp only [hx, Bool.if_false_right,
Bool.false_eq_true, ↓reduceIte, Nat.zero_add, testBit_toNat]
have hi : i < w ∨ (w ≤ i ∧ i < w + k) ∨ w + k ≤ i := by omega
rcases hi with hi | hi | hi
· simp [hi]; omega
Expand Down

0 comments on commit 7692343

Please sign in to comment.