Skip to content

Commit edeae18

Browse files
luisacicolinialexkeizerkim-em
authored
feat: add Bitvec reverse definition, getLsbD_reverse, getMsbD_reverse, reverse_append, reverse_replicate and Nat.mod_sub_eq_sub_mod (#6476)
This PR defines `reverse` for bitvectors and implements a first subset of theorems (`getLsbD_reverse, getMsbD_reverse, reverse_append, reverse_replicate, reverse_cast, msb_reverse`). We also include some necessary related theorems (`cons_append, cons_append_append, append_assoc, replicate_append_self, replicate_succ'`) and deprecate theorems`replicate_zero_eq` and `replicate_succ_eq`. --------- Co-authored-by: Alex Keizer <alex@keizer.dev> Co-authored-by: Kim Morrison <kim@tqft.net>
1 parent 91bae2e commit edeae18

File tree

2 files changed

+120
-3
lines changed

2 files changed

+120
-3
lines changed

src/Init/Data/BitVec/Basic.lean

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -669,4 +669,11 @@ def ofBoolListLE : (bs : List Bool) → BitVec bs.length
669669
| [] => 0#0
670670
| b :: bs => concat (ofBoolListLE bs) b
671671

672+
/- ### reverse -/
673+
674+
/-- Reverse the bits in a bitvector. -/
675+
def reverse : {w : Nat} → BitVec w → BitVec w
676+
| 0, x => x
677+
| w + 1, x => concat (reverse (x.truncate w)) (x.msb)
678+
672679
end BitVec

src/Init/Data/BitVec/Lemmas.lean

Lines changed: 113 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -2053,6 +2053,32 @@ theorem eq_msb_cons_setWidth (x : BitVec (w+1)) : x = (cons x.msb (x.setWidth w)
20532053
ext i
20542054
simp [cons]
20552055

2056+
2057+
theorem cons_append (x : BitVec w₁) (y : BitVec w₂) (a : Bool) :
2058+
(cons a x) ++ y = (cons a (x ++ y)).cast (by omega) := by
2059+
apply eq_of_toNat_eq
2060+
simp only [toNat_append, toNat_cons, toNat_cast]
2061+
rw [Nat.shiftLeft_add, Nat.shiftLeft_or_distrib, Nat.or_assoc]
2062+
2063+
theorem cons_append_append (x : BitVec w₁) (y : BitVec w₂) (z : BitVec w₃) (a : Bool) :
2064+
(cons a x) ++ y ++ z = (cons a (x ++ y ++ z)).cast (by omega) := by
2065+
ext i h
2066+
simp only [cons, getLsbD_append, getLsbD_cast, getLsbD_ofBool, cast_cast]
2067+
by_cases h₀ : i < w₁ + w₂ + w₃
2068+
· simp only [h₀, ↓reduceIte]
2069+
by_cases h₁ : i < w₃
2070+
· simp [h₁]
2071+
· simp only [h₁, ↓reduceIte]
2072+
by_cases h₂ : i - w₃ < w₂
2073+
· simp [h₂]
2074+
· simp [h₂]
2075+
omega
2076+
· simp only [show ¬i - w₃ - w₂ < w₁ by omega, ↓reduceIte, show i - w₃ - w₂ - w₁ = 0 by omega,
2077+
decide_true, Bool.true_and, h₀, show i - (w₁ + w₂ + w₃) = 0 by omega]
2078+
by_cases h₂ : i < w₃
2079+
· simp [h₂]; omega
2080+
· simp [h₂]; omega
2081+
20562082
/-! ### concat -/
20572083

20582084
@[simp] theorem toNat_concat (x : BitVec w) (b : Bool) :
@@ -3373,11 +3399,11 @@ theorem and_one_eq_setWidth_ofBool_getLsbD {x : BitVec w} :
33733399
ext (_ | i) h <;> simp [Bool.and_comm]
33743400

33753401
@[simp]
3376-
theorem replicate_zero_eq {x : BitVec w} : x.replicate 0 = 0#0 := by
3402+
theorem replicate_zero {x : BitVec w} : x.replicate 0 = 0#0 := by
33773403
simp [replicate]
33783404

33793405
@[simp]
3380-
theorem replicate_succ_eq {x : BitVec w} :
3406+
theorem replicate_succ {x : BitVec w} :
33813407
x.replicate (n + 1) =
33823408
(x ++ replicate n x).cast (by rw [Nat.mul_succ]; omega) := by
33833409
simp [replicate]
@@ -3389,7 +3415,7 @@ theorem getLsbD_replicate {n w : Nat} (x : BitVec w) :
33893415
induction n generalizing x
33903416
case zero => simp
33913417
case succ n ih =>
3392-
simp only [replicate_succ_eq, getLsbD_cast, getLsbD_append]
3418+
simp only [replicate_succ, getLsbD_cast, getLsbD_append]
33933419
by_cases hi : i < w * (n + 1)
33943420
· simp only [hi, decide_true, Bool.true_and]
33953421
by_cases hi' : i < w * n
@@ -3406,6 +3432,33 @@ theorem getElem_replicate {n w : Nat} (x : BitVec w) (h : i < w * n) :
34063432
simp only [← getLsbD_eq_getElem, getLsbD_replicate]
34073433
by_cases h' : w = 0 <;> simp [h'] <;> omega
34083434

3435+
theorem append_assoc {x₁ : BitVec w₁} {x₂ : BitVec w₂} {x₃ : BitVec w₃} :
3436+
(x₁ ++ x₂) ++ x₃ = (x₁ ++ (x₂ ++ x₃)).cast (by omega) := by
3437+
induction w₁ generalizing x₂ x₃
3438+
case zero => simp
3439+
case succ n ih =>
3440+
specialize @ih (setWidth n x₁)
3441+
rw [← cons_msb_setWidth x₁, cons_append_append, ih, cons_append]
3442+
ext j h
3443+
simp [getLsbD_cons, show n + w₂ + w₃ = n + (w₂ + w₃) by omega]
3444+
3445+
theorem replicate_append_self {x : BitVec w} :
3446+
x ++ x.replicate n = (x.replicate n ++ x).cast (by omega) := by
3447+
induction n with
3448+
| zero => simp
3449+
| succ n ih =>
3450+
rw [replicate_succ]
3451+
conv => lhs; rw [ih]
3452+
simp only [cast_cast, cast_eq]
3453+
rw [← cast_append_left]
3454+
· rw [append_assoc]; congr
3455+
· rw [Nat.add_comm, Nat.mul_add, Nat.mul_one]; omega
3456+
3457+
theorem replicate_succ' {x : BitVec w} :
3458+
x.replicate (n + 1) =
3459+
(replicate n x ++ x).cast (by rw [Nat.mul_succ]) := by
3460+
simp [replicate_append_self]
3461+
34093462
/-! ### intMin -/
34103463

34113464
/-- The bitvector of width `w` that has the smallest value when interpreted as an integer. -/
@@ -3691,6 +3744,57 @@ theorem toInt_abs_eq_natAbs_of_ne_intMin {x : BitVec w} (hx : x ≠ intMin w) :
36913744
x.abs.toInt = x.toInt.natAbs := by
36923745
simp [toInt_abs_eq_natAbs, hx]
36933746

3747+
/-! ### Reverse -/
3748+
3749+
theorem getLsbD_reverse {i : Nat} {x : BitVec w} :
3750+
(x.reverse).getLsbD i = x.getMsbD i := by
3751+
induction w generalizing i
3752+
case zero => simp
3753+
case succ n ih =>
3754+
simp only [reverse, truncate_eq_setWidth, getLsbD_concat]
3755+
rcases i with rfl | i
3756+
· rfl
3757+
· simp only [Nat.add_one_ne_zero, ↓reduceIte, Nat.add_one_sub_one, ih]
3758+
rw [getMsbD_setWidth]
3759+
simp only [show n - (n + 1) = 0 by omega, Nat.zero_le, decide_true, Bool.true_and]
3760+
congr; omega
3761+
3762+
theorem getMsbD_reverse {i : Nat} {x : BitVec w} :
3763+
(x.reverse).getMsbD i = x.getLsbD i := by
3764+
simp only [getMsbD_eq_getLsbD, getLsbD_reverse]
3765+
by_cases hi : i < w
3766+
· simp only [hi, decide_true, show w - 1 - i < w by omega, Bool.true_and]
3767+
congr; omega
3768+
· simp [hi, show i ≥ w by omega]
3769+
3770+
theorem msb_reverse {x : BitVec w} :
3771+
(x.reverse).msb = x.getLsbD 0 :=
3772+
by rw [BitVec.msb, getMsbD_reverse]
3773+
3774+
theorem reverse_append {x : BitVec w} {y : BitVec v} :
3775+
(x ++ y).reverse = (y.reverse ++ x.reverse).cast (by omega) := by
3776+
ext i h
3777+
simp only [getLsbD_append, getLsbD_reverse]
3778+
by_cases hi : i < v
3779+
· by_cases hw : w ≤ i
3780+
· simp [getMsbD_append, getLsbD_cast, getLsbD_append, getLsbD_reverse, hw]
3781+
· simp [getMsbD_append, getLsbD_cast, getLsbD_append, getLsbD_reverse, hw, show i < w by omega]
3782+
· by_cases hw : w ≤ i
3783+
· simp [getMsbD_append, getLsbD_cast, getLsbD_append, hw, show ¬ i < w by omega, getLsbD_reverse]
3784+
· simp [getMsbD_append, getLsbD_cast, getLsbD_append, hw, show i < w by omega, getLsbD_reverse]
3785+
3786+
@[simp]
3787+
theorem reverse_cast {w v : Nat} (h : w = v) (x : BitVec w) :
3788+
(x.cast h).reverse = x.reverse.cast h := by
3789+
subst h; simp
3790+
3791+
theorem reverse_replicate {n : Nat} {x : BitVec w} :
3792+
(x.replicate n).reverse = (x.reverse).replicate n := by
3793+
induction n with
3794+
| zero => rfl
3795+
| succ n ih =>
3796+
conv => lhs; simp only [replicate_succ']
3797+
simp [reverse_append, ih]
36943798

36953799
/-! ### Decidable quantifiers -/
36963800

@@ -3906,4 +4010,10 @@ abbrev shiftLeft_zero_eq := @shiftLeft_zero
39064010
@[deprecated ushiftRight_zero (since := "2024-10-27")]
39074011
abbrev ushiftRight_zero_eq := @ushiftRight_zero
39084012

4013+
@[deprecated replicate_zero (since := "2025-01-08")]
4014+
abbrev replicate_zero_eq := @replicate_zero
4015+
4016+
@[deprecated replicate_succ (since := "2025-01-08")]
4017+
abbrev replicate_succ_eq := @replicate_succ
4018+
39094019
end BitVec

0 commit comments

Comments
 (0)