Skip to content

Commit

Permalink
feat: BitVec.[toInt|toFin|getMsbD]_ofBool (#6317)
Browse files Browse the repository at this point in the history
This PR completes the basic API for BitVec.ofBool.

---------

Co-authored-by: Kim Morrison <[email protected]>
  • Loading branch information
tobiasgrosser and kim-em authored Dec 10, 2024
1 parent a805946 commit 1786539
Showing 1 changed file with 11 additions and 2 deletions.
13 changes: 11 additions & 2 deletions src/Init/Data/BitVec/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -241,8 +241,11 @@ theorem toFin_one : toFin (1 : BitVec w) = 1 := by
@[simp] theorem toNat_ofBool (b : Bool) : (ofBool b).toNat = b.toNat := by
cases b <;> rfl

@[simp] theorem msb_ofBool (b : Bool) : (ofBool b).msb = b := by
cases b <;> simp [BitVec.msb, getMsbD, getLsbD]
@[simp] theorem toInt_ofBool (b : Bool) : (ofBool b).toInt = -b.toInt := by
cases b <;> rfl

@[simp] theorem toFin_ofBool (b : Bool) : (ofBool b).toFin = Fin.ofNat' 2 (b.toNat) := by
cases b <;> rfl

theorem ofNat_one (n : Nat) : BitVec.ofNat 1 n = BitVec.ofBool (n % 2 = 1) := by
rcases (Nat.mod_two_eq_zero_or_one n) with h | h <;> simp [h, BitVec.ofNat, Fin.ofNat']
Expand Down Expand Up @@ -366,6 +369,12 @@ theorem getElem_ofBool {b : Bool} {i : Nat} : (ofBool b)[0] = b := by
· simp only [ofBool]
by_cases hi : i = 0 <;> simp [hi] <;> omega

@[simp] theorem getMsbD_ofBool (b : Bool) : (ofBool b).getMsbD i = (decide (i = 0) && b) := by
cases b <;> simp [getMsbD]

@[simp] theorem msb_ofBool (b : Bool) : (ofBool b).msb = b := by
cases b <;> simp [BitVec.msb]

/-! ### msb -/

@[simp] theorem msb_zero : (0#w).msb = false := by simp [BitVec.msb, getMsbD]
Expand Down

0 comments on commit 1786539

Please sign in to comment.