Skip to content

Commit 971eede

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 acde022 commit 971eede

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
@@ -3711,14 +3711,19 @@ theorem and_one_eq_setWidth_ofBool_getLsbD {x : BitVec w} :
37113711
theorem replicate_zero {x : BitVec w} : x.replicate 0 = 0#0 := by
37123712
simp [replicate]
37133713

3714+
@[simp]
3715+
theorem replicate_one {w : Nat} {x : BitVec w} :
3716+
(x.replicate 1) = x.cast (by rw [Nat.mul_one]) := by
3717+
simp [replicate]
3718+
37143719
@[simp]
37153720
theorem replicate_succ {x : BitVec w} :
37163721
x.replicate (n + 1) =
37173722
(x ++ replicate n x).cast (by rw [Nat.mul_succ]; omega) := by
37183723
simp [replicate]
37193724

37203725
@[simp]
3721-
theorem getLsbD_replicate {n w : Nat} (x : BitVec w) :
3726+
theorem getLsbD_replicate {n w : Nat} {x : BitVec w} :
37223727
(x.replicate n).getLsbD i =
37233728
(decide (i < w * n) && x.getLsbD (i % w)) := by
37243729
induction n generalizing x
@@ -3729,17 +3734,17 @@ theorem getLsbD_replicate {n w : Nat} (x : BitVec w) :
37293734
· simp only [hi, decide_true, Bool.true_and]
37303735
by_cases hi' : i < w * n
37313736
· simp [hi', ih]
3732-
· simp [hi', decide_false]
3733-
rw [Nat.sub_mul_eq_mod_of_lt_of_le] <;> omega
3737+
· simp only [hi', ↓reduceIte]
3738+
rw [Nat.sub_mul_eq_mod_of_lt_of_le (by omega) (by omega)]
37343739
· rw [Nat.mul_succ] at hi ⊢
37353740
simp only [show ¬i < w * n by omega, decide_false, cond_false, hi, Bool.false_and]
37363741
apply BitVec.getLsbD_ge (x := x) (i := i - w * n) (ge := by omega)
37373742

37383743
@[simp]
3739-
theorem getElem_replicate {n w : Nat} (x : BitVec w) (h : i < w * n) :
3744+
theorem getElem_replicate {n w : Nat} {x : BitVec w} (h : i < w * n) :
37403745
(x.replicate n)[i] = if h' : w = 0 then false else x[i % w]'(@Nat.mod_lt i w (by omega)) := by
37413746
simp only [← getLsbD_eq_getElem, getLsbD_replicate]
3742-
by_cases h' : w = 0 <;> simp [h'] <;> omega
3747+
cases w <;> simp; omega
37433748

37443749
theorem append_assoc {x₁ : BitVec w₁} {x₂ : BitVec w₂} {x₃ : BitVec w₃} :
37453750
(x₁ ++ x₂) ++ x₃ = (x₁ ++ (x₂ ++ x₃)).cast (by omega) := by
@@ -4113,6 +4118,17 @@ theorem reverse_replicate {n : Nat} {x : BitVec w} :
41134118
conv => lhs; simp only [replicate_succ']
41144119
simp [reverse_append, ih]
41154120

4121+
@[simp]
4122+
theorem getMsbD_replicate {n w : Nat} {x : BitVec w} :
4123+
(x.replicate n).getMsbD i = (decide (i < w * n) && x.getMsbD (i % w)) := by
4124+
rw [← getLsbD_reverse, reverse_replicate, getLsbD_replicate, getLsbD_reverse]
4125+
4126+
@[simp]
4127+
theorem msb_replicate {n w : Nat} {x : BitVec w} :
4128+
(x.replicate n).msb = (decide (0 < n) && x.msb) := by
4129+
simp only [BitVec.msb, getMsbD_replicate, Nat.zero_mod]
4130+
cases n <;> cases w <;> simp
4131+
41164132
/-! ### Decidable quantifiers -/
41174133

41184134
theorem forall_zero_iff {P : BitVec 0Prop} :

0 commit comments

Comments
 (0)