Skip to content

Commit

Permalink
chore: move theorems around to proper location
Browse files Browse the repository at this point in the history
  • Loading branch information
bollu committed Jun 8, 2024
1 parent 8632849 commit 351a7ce
Showing 1 changed file with 38 additions and 37 deletions.
75 changes: 38 additions & 37 deletions src/Init/Data/BitVec/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -163,6 +163,16 @@ theorem toNat_zero (n : Nat) : (0#n).toNat = 0 := by trivial
private theorem lt_two_pow_of_le {x m n : Nat} (lt : x < 2 ^ m) (le : m ≤ n) : x < 2 ^ n :=
Nat.lt_of_lt_of_le lt (Nat.pow_le_pow_of_le_right (by trivial : 0 < 2) le)

@[simp]
theorem getLsb_ofBool (b : Bool) (i : Nat) : (BitVec.ofBool b).getLsb i = ((i = 0) && b) := by
rcases b with rfl | rfl
· simp [ofBool]
· simp [ofBool, getLsb_ofNat]
by_cases hi : (i = 0)
· simp [hi]
· simp [hi]
omega

/-! ### msb -/

@[simp] theorem msb_zero : (0#w).msb = false := by simp [BitVec.msb, getMsb]
Expand Down Expand Up @@ -408,6 +418,29 @@ theorem msb_zeroExtend (x : BitVec w) : (x.zeroExtend v).msb = (decide (0 < v) &
theorem msb_zeroExtend' (x : BitVec w) (h : w ≤ v) : (x.zeroExtend' h).msb = (decide (0 < v) && x.getLsb (v - 1)) := by
rw [zeroExtend'_eq, msb_zeroExtend]

/-- zero extending a bitvector to width 1 equals the boolean of the lsb. -/
theorem zeroExtend_one_eq_ofBool_getLsb_zero (x : BitVec w) :
x.zeroExtend 1 = BitVec.ofBool (x.getLsb 0) := by
ext i
simp [getLsb_zeroExtend, Fin.fin_one_eq_zero i]

/-- `testBit 1 i` is true iff the index `i` equals 0. -/
private theorem Nat.testBit_one_eq_true_iff_self_eq_zero {i : Nat} :
Nat.testBit 1 i = true ↔ i = 0 := by
cases i <;> simp

/-- Zero extending `1#v` to `1#w` equals `1#w` when `v > 0`. -/
theorem zeroExtend_ofNat_one_eq_ofNat_one_of_lt {v w : Nat} (hv : 0 < v):
(BitVec.ofNat v 1).zeroExtend w = BitVec.ofNat w 1 := by
ext i
obtain ⟨i, hilt⟩ := i
simp only [getLsb_zeroExtend, hilt, decide_True, getLsb_ofNat, Bool.true_and,
Bool.and_iff_right_iff_imp, decide_eq_true_eq]
intros hi1
have hv := Nat.testBit_one_eq_true_iff_self_eq_zero.mp hi1
omega


/-! ## extractLsb -/

@[simp]
Expand Down Expand Up @@ -593,6 +626,11 @@ theorem not_def {x : BitVec v} : ~~~x = allOnes v ^^^ x := rfl
@[simp] theorem toFin_shiftLeft {n : Nat} (x : BitVec w) :
BitVec.toFin (x <<< n) = Fin.ofNat' (x.toNat <<< n) (Nat.two_pow_pos w) := rfl

@[simp]
theorem shiftLeft_zero_eq (x : BitVec w) : x <<< 0 = x := by
apply eq_of_toNat_eq
simp

@[simp] theorem getLsb_shiftLeft (x : BitVec m) (n) :
getLsb (x <<< n) i = (decide (i < m) && !decide (i < n) && getLsb x (i - n)) := by
rw [← testBit_toNat, getLsb]
Expand Down Expand Up @@ -1100,11 +1138,6 @@ def mulRec (l r : BitVec w) (s : Nat) : BitVec w :=
| 0 => cur
| s + 1 => mulRec l r s + cur

@[simp]
theorem shiftLeft_zero_eq (x : BitVec w) : x <<< 0 = x := by
apply eq_of_toNat_eq
simp

theorem mulRec_zero_eq (l r : BitVec w) :
mulRec l r 0 = if r.getLsb 0 then l else 0 := by
simp [mulRec]
Expand All @@ -1113,38 +1146,6 @@ theorem mulRec_succ_eq (l r : BitVec w) (s : Nat) :
mulRec l r (s + 1) = mulRec l r s + if r.getLsb (s + 1) then (l <<< (s + 1)) else 0 := by
simp [mulRec]

@[simp]
theorem getLsb_ofBool (b : Bool) (i : Nat) : (BitVec.ofBool b).getLsb i = ((i = 0) && b) := by
rcases b with rfl | rfl
· simp [ofBool]
· simp [ofBool, getLsb_ofNat]
by_cases hi : (i = 0)
· simp [hi]
· simp [hi]
omega

/-- zero extending a bitvector to width 1 equals the boolean of the lsb. -/
theorem zeroExtend_one_eq_ofBool_getLsb_zero (x : BitVec w) :
x.zeroExtend 1 = BitVec.ofBool (x.getLsb 0) := by
ext i
simp [getLsb_zeroExtend, Fin.fin_one_eq_zero i]

/-- `testBit 1 i` is true iff the index `i` equals 0. -/
private theorem Nat.testBit_one_eq_true_iff_self_eq_zero {i : Nat} :
Nat.testBit 1 i = true ↔ i = 0 := by
cases i <;> simp

/-- Zero extending `1#v` to `1#w` equals `1#w` when `v > 0`. -/
theorem zeroExtend_ofNat_one_eq_ofNat_one_of_lt {v w : Nat} (hv : 0 < v):
(BitVec.ofNat v 1).zeroExtend w = BitVec.ofNat w 1 := by
ext i
obtain ⟨i, hilt⟩ := i
simp only [getLsb_zeroExtend, hilt, decide_True, getLsb_ofNat, Bool.true_and,
Bool.and_iff_right_iff_imp, decide_eq_true_eq]
intros hi1
have hv := Nat.testBit_one_eq_true_iff_self_eq_zero.mp hi1
omega

@[simp]
theorem BitVec.mul_one {x : BitVec w} : x * (1#w) = x := by
apply eq_of_toNat_eq
Expand Down

0 comments on commit 351a7ce

Please sign in to comment.