Skip to content

Commit

Permalink
feat: getLsb_replicate
Browse files Browse the repository at this point in the history
This allows bitblasting `BitVec.replicate`.
  • Loading branch information
bollu committed Jul 30, 2024
1 parent 92cca5e commit a82aebc
Show file tree
Hide file tree
Showing 2 changed files with 43 additions and 4 deletions.
6 changes: 2 additions & 4 deletions src/Init/Data/BitVec/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -583,11 +583,9 @@ instance : HAppend (BitVec w) (BitVec v) (BitVec (w + v)) := ⟨.append⟩
-- TODO: write this using multiplication
/-- `replicate i x` concatenates `i` copies of `x` into a new vector of length `w*i`. -/
def replicate : (i : Nat) → BitVec w → BitVec (w*i)
| 0, _ => 0
| 0, _ => 0#0
| n+1, x =>
have hEq : w + w*n = w*(n + 1) := by
rw [Nat.mul_add, Nat.add_comm, Nat.mul_one]
hEq ▸ (x ++ replicate n x)
(x ++ replicate n x).cast (by rw [Nat.mul_succ]; omega)

/-!
### Cons and Concat
Expand Down
41 changes: 41 additions & 0 deletions src/Init/Data/BitVec/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1504,4 +1504,45 @@ theorem zeroExtend_truncate_succ_eq_zeroExtend_truncate_or_twoPow_of_getLsb_true
simp [hx]
· by_cases hik' : k < i + 1 <;> simp [hik, hik'] <;> omega

@[simp]
theorem replicate_zero_eq {x : BitVec w} : x.replicate 0 = 0#0 := by
simp [replicate]

@[simp]
theorem replicate_succ_eq {x : BitVec w} :
x.replicate (n + 1) =
(x ++ replicate n x).cast (by rw [Nat.mul_succ]; omega) := by
simp [replicate]

/--
If a number `w * n ≤ i < w * (n + 1)`, then `i - w * n` equals `i % w`.
Intuitively, this is true by subtracting `w * n` from the inequality, giving
`0 ≤ i - w * n < w`, which uniquely identifies `i % w`.
-/
private theorem Nat.sub_mul_eq_mod_of_lt_of_le (hlo : w * n ≤ i) (hhi : i < w * (n + 1)) :
i - w * n = i % w := by
rw [Nat.mul_succ] at hhi
rw [Nat.sub_eq_of_eq_add]
suffices i / w = n by rw [← this, Nat.add_comm, Nat.div_add_mod]
apply Nat.div_eq_of_lt_le
(by rw [Nat.mul_comm]; omega)
(by rw [Nat.mul_comm, Nat.mul_succ]; omega)

theorem getLsb_replicate {n w : Nat} (x : BitVec w) :
(x.replicate n).getLsb i =
((decide (i < w * n)) && (x.getLsb (i % w))) := by
induction n generalizing x
case zero => simp
case succ n ih =>
simp only [replicate_succ_eq, getLsb_cast, getLsb_append]
by_cases hi : i < w * (n + 1)
· simp only [hi, decide_True, Bool.true_and]
by_cases hi' : i < w * n
· simp [hi', ih]
· simp only [hi', decide_False, cond_false]
rw [Nat.sub_mul_eq_mod_of_lt_of_le] <;> omega
· rw [Nat.mul_succ] at hi ⊢
simp only [show ¬i < w * n by omega, decide_False, cond_false, hi, Bool.false_and]
apply BitVec.getLsb_ge (x := x) (i := i - w * n) (ge := by omega)

end BitVec

0 comments on commit a82aebc

Please sign in to comment.