Skip to content

Commit

Permalink
feat: BitVec.toInt_shiftLeft theorem (#6346)
Browse files Browse the repository at this point in the history
This PR completes the toNat/Int/Fin family for `shiftLeft`.
  • Loading branch information
mhk119 authored Jan 8, 2025
1 parent 18b183f commit 91cbd7c
Showing 1 changed file with 7 additions and 2 deletions.
9 changes: 7 additions & 2 deletions src/Init/Data/BitVec/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1142,11 +1142,16 @@ theorem getMsb_not {x : BitVec w} :
/-! ### shiftLeft -/

@[simp, bv_toNat] theorem toNat_shiftLeft {x : BitVec v} :
BitVec.toNat (x <<< n) = BitVec.toNat x <<< n % 2^v :=
(x <<< n).toNat = x.toNat <<< n % 2^v :=
BitVec.toNat_ofNat _ _

@[simp] theorem toInt_shiftLeft {x : BitVec w} :
(x <<< n).toInt = (x.toNat <<< n : Int).bmod (2^w) := by
rw [toInt_eq_toNat_bmod, toNat_shiftLeft, Nat.shiftLeft_eq]
simp

@[simp] theorem toFin_shiftLeft {n : Nat} (x : BitVec w) :
BitVec.toFin (x <<< n) = Fin.ofNat' (2^w) (x.toNat <<< n) := rfl
(x <<< n).toFin = Fin.ofNat' (2^w) (x.toNat <<< n) := rfl

@[simp]
theorem shiftLeft_zero (x : BitVec w) : x <<< 0 = x := by
Expand Down

0 comments on commit 91cbd7c

Please sign in to comment.