Skip to content

Commit 7cceaf8

Browse files
feat: add BitVec.(getMsbD, msb)_replicate, replicate_one (leanprover#6326)
This PR adds `BitVec.(getMsbD, msb)_replicate, replicate_one` theorems, corrects a non-terminal `simp` in `BitVec.getLsbD_replicate` and simplifies the proof of `BitVec.getElem_replicate` using the `cases` tactic. Co-authored with @bollu. --------- Co-authored-by: Alex Keizer <[email protected]>
1 parent 5b44298 commit 7cceaf8

File tree

1 file changed

+21
-5
lines changed

1 file changed

+21
-5
lines changed

src/Init/Data/BitVec/Lemmas.lean

+21-5
Original file line numberDiff line numberDiff line change
@@ -3732,14 +3732,19 @@ theorem and_one_eq_setWidth_ofBool_getLsbD {x : BitVec w} :
37323732
theorem replicate_zero {x : BitVec w} : x.replicate 0 = 0#0 := by
37333733
simp [replicate]
37343734

3735+
@[simp]
3736+
theorem replicate_one {w : Nat} {x : BitVec w} :
3737+
(x.replicate 1) = x.cast (by rw [Nat.mul_one]) := by
3738+
simp [replicate]
3739+
37353740
@[simp]
37363741
theorem replicate_succ {x : BitVec w} :
37373742
x.replicate (n + 1) =
37383743
(x ++ replicate n x).cast (by rw [Nat.mul_succ]; omega) := by
37393744
simp [replicate]
37403745

37413746
@[simp]
3742-
theorem getLsbD_replicate {n w : Nat} (x : BitVec w) :
3747+
theorem getLsbD_replicate {n w : Nat} {x : BitVec w} :
37433748
(x.replicate n).getLsbD i =
37443749
(decide (i < w * n) && x.getLsbD (i % w)) := by
37453750
induction n generalizing x
@@ -3750,17 +3755,17 @@ theorem getLsbD_replicate {n w : Nat} (x : BitVec w) :
37503755
· simp only [hi, decide_true, Bool.true_and]
37513756
by_cases hi' : i < w * n
37523757
· simp [hi', ih]
3753-
· simp [hi', decide_false]
3754-
rw [Nat.sub_mul_eq_mod_of_lt_of_le] <;> omega
3758+
· simp only [hi', ↓reduceIte]
3759+
rw [Nat.sub_mul_eq_mod_of_lt_of_le (by omega) (by omega)]
37553760
· rw [Nat.mul_succ] at hi ⊢
37563761
simp only [show ¬i < w * n by omega, decide_false, cond_false, hi, Bool.false_and]
37573762
apply BitVec.getLsbD_ge (x := x) (i := i - w * n) (ge := by omega)
37583763

37593764
@[simp]
3760-
theorem getElem_replicate {n w : Nat} (x : BitVec w) (h : i < w * n) :
3765+
theorem getElem_replicate {n w : Nat} {x : BitVec w} (h : i < w * n) :
37613766
(x.replicate n)[i] = if h' : w = 0 then false else x[i % w]'(@Nat.mod_lt i w (by omega)) := by
37623767
simp only [← getLsbD_eq_getElem, getLsbD_replicate]
3763-
by_cases h' : w = 0 <;> simp [h'] <;> omega
3768+
cases w <;> simp; omega
37643769

37653770
theorem append_assoc {x₁ : BitVec w₁} {x₂ : BitVec w₂} {x₃ : BitVec w₃} :
37663771
(x₁ ++ x₂) ++ x₃ = (x₁ ++ (x₂ ++ x₃)).cast (by omega) := by
@@ -4126,6 +4131,17 @@ theorem reverse_replicate {n : Nat} {x : BitVec w} :
41264131
conv => lhs; simp only [replicate_succ']
41274132
simp [reverse_append, ih]
41284133

4134+
@[simp]
4135+
theorem getMsbD_replicate {n w : Nat} {x : BitVec w} :
4136+
(x.replicate n).getMsbD i = (decide (i < w * n) && x.getMsbD (i % w)) := by
4137+
rw [← getLsbD_reverse, reverse_replicate, getLsbD_replicate, getLsbD_reverse]
4138+
4139+
@[simp]
4140+
theorem msb_replicate {n w : Nat} {x : BitVec w} :
4141+
(x.replicate n).msb = (decide (0 < n) && x.msb) := by
4142+
simp only [BitVec.msb, getMsbD_replicate, Nat.zero_mod]
4143+
cases n <;> cases w <;> simp
4144+
41294145
/-! ### Decidable quantifiers -/
41304146

41314147
theorem forall_zero_iff {P : BitVec 0Prop} :

0 commit comments

Comments
 (0)