Skip to content

Commit e106be1

Browse files
bollutobiasgrosser
andauthored
feat: sshiftRight bitblasting (#4889)
We follow the same strategy as #4872, #4571, and implement bitblasting theorems for `sshiftRight`. --------- Co-authored-by: Tobias Grosser <tobias@grosser.es>
1 parent 1efd665 commit e106be1

File tree

3 files changed

+106
-1
lines changed

3 files changed

+106
-1
lines changed

src/Init/Data/BitVec/Basic.lean

Lines changed: 9 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -536,6 +536,15 @@ def sshiftRight (a : BitVec n) (s : Nat) : BitVec n := .ofInt n (a.toInt >>> s)
536536
instance {n} : HShiftLeft (BitVec m) (BitVec n) (BitVec m) := ⟨fun x y => x <<< y.toNat⟩
537537
instance {n} : HShiftRight (BitVec m) (BitVec n) (BitVec m) := ⟨fun x y => x >>> y.toNat⟩
538538

539+
/--
540+
Arithmetic right shift for bit vectors. The high bits are filled with the
541+
most-significant bit.
542+
As a numeric operation, this is equivalent to `a.toInt >>> s.toNat`.
543+
544+
SMT-Lib name: `bvashr`.
545+
-/
546+
def sshiftRight' (a : BitVec n) (s : BitVec m) : BitVec n := a.sshiftRight s.toNat
547+
539548
/-- Auxiliary function for `rotateLeft`, which does not take into account the case where
540549
the rotation amount is greater than the bitvector width. -/
541550
def rotateLeftAux (x : BitVec w) (n : Nat) : BitVec w :=

src/Init/Data/BitVec/Bitblast.lean

Lines changed: 61 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -429,6 +429,67 @@ theorem shiftLeft_eq_shiftLeftRec (x : BitVec w₁) (y : BitVec w₂) :
429429
· simp [of_length_zero]
430430
· simp [shiftLeftRec_eq]
431431

432+
/- ### Arithmetic shift right (sshiftRight) recurrence -/
433+
434+
/--
435+
`sshiftRightRec x y n` shifts `x` arithmetically/signed to the right by the first `n` bits of `y`.
436+
The theorem `sshiftRight_eq_sshiftRightRec` proves the equivalence of `(x.sshiftRight y)` and `sshiftRightRec`.
437+
Together with equations `sshiftRightRec_zero`, `sshiftRightRec_succ`,
438+
this allows us to unfold `sshiftRight` into a circuit for bitblasting.
439+
-/
440+
def sshiftRightRec (x : BitVec w₁) (y : BitVec w₂) (n : Nat) : BitVec w₁ :=
441+
let shiftAmt := (y &&& (twoPow w₂ n))
442+
match n with
443+
| 0 => x.sshiftRight' shiftAmt
444+
| n + 1 => (sshiftRightRec x y n).sshiftRight' shiftAmt
445+
446+
@[simp]
447+
theorem sshiftRightRec_zero_eq (x : BitVec w₁) (y : BitVec w₂) :
448+
sshiftRightRec x y 0 = x.sshiftRight' (y &&& 1#w₂) := by
449+
simp only [sshiftRightRec, twoPow_zero]
450+
451+
@[simp]
452+
theorem sshiftRightRec_succ_eq (x : BitVec w₁) (y : BitVec w₂) (n : Nat) :
453+
sshiftRightRec x y (n + 1) = (sshiftRightRec x y n).sshiftRight' (y &&& twoPow w₂ (n + 1)) := by
454+
simp [sshiftRightRec]
455+
456+
/--
457+
If `y &&& z = 0`, `x.sshiftRight (y ||| z) = (x.sshiftRight y).sshiftRight z`.
458+
This follows as `y &&& z = 0` implies `y ||| z = y + z`,
459+
and thus `x.sshiftRight (y ||| z) = x.sshiftRight (y + z) = (x.sshiftRight y).sshiftRight z`.
460+
-/
461+
theorem sshiftRight'_or_of_and_eq_zero {x : BitVec w₁} {y z : BitVec w₂}
462+
(h : y &&& z = 0#w₂) :
463+
x.sshiftRight' (y ||| z) = (x.sshiftRight' y).sshiftRight' z := by
464+
simp [sshiftRight', ← add_eq_or_of_and_eq_zero _ _ h,
465+
toNat_add_of_and_eq_zero h, sshiftRight_add]
466+
467+
theorem sshiftRightRec_eq (x : BitVec w₁) (y : BitVec w₂) (n : Nat) :
468+
sshiftRightRec x y n = x.sshiftRight' ((y.truncate (n + 1)).zeroExtend w₂) := by
469+
induction n generalizing x y
470+
case zero =>
471+
ext i
472+
simp [twoPow_zero, Nat.reduceAdd, and_one_eq_zeroExtend_ofBool_getLsb, truncate_one]
473+
case succ n ih =>
474+
simp only [sshiftRightRec_succ_eq, and_twoPow, ih]
475+
by_cases h : y.getLsb (n + 1)
476+
· rw [zeroExtend_truncate_succ_eq_zeroExtend_truncate_or_twoPow_of_getLsb_true h,
477+
sshiftRight'_or_of_and_eq_zero (by simp), h]
478+
simp
479+
· rw [zeroExtend_truncate_succ_eq_zeroExtend_truncate_of_getLsb_false (i := n + 1)
480+
(by simp [h])]
481+
simp [h]
482+
483+
/--
484+
Show that `x.sshiftRight y` can be written in terms of `sshiftRightRec`.
485+
This can be unfolded in terms of `sshiftRightRec_zero_eq`, `sshiftRightRec_succ_eq` for bitblasting.
486+
-/
487+
theorem sshiftRight_eq_sshiftRightRec (x : BitVec w₁) (y : BitVec w₂) :
488+
(x.sshiftRight' y).getLsb i = (sshiftRightRec x y (w₂ - 1)).getLsb i := by
489+
rcases w₂ with rfl | w₂
490+
· simp [of_length_zero]
491+
· simp [sshiftRightRec_eq]
492+
432493
/- ### Logical shift right (ushiftRight) recurrence for bitblasting -/
433494

434495
/--

src/Init/Data/BitVec/Lemmas.lean

Lines changed: 36 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -786,7 +786,7 @@ theorem sshiftRight_eq_of_msb_true {x : BitVec w} {s : Nat} (h : x.msb = true) :
786786
· rw [Nat.shiftRight_eq_div_pow]
787787
apply Nat.lt_of_le_of_lt (Nat.div_le_self _ _) (by omega)
788788

789-
theorem getLsb_sshiftRight (x : BitVec w) (s i : Nat) :
789+
@[simp] theorem getLsb_sshiftRight (x : BitVec w) (s i : Nat) :
790790
getLsb (x.sshiftRight s) i =
791791
(!decide (w ≤ i) && if s + i < w then x.getLsb (s + i) else x.msb) := by
792792
rcases hmsb : x.msb with rfl | rfl
@@ -807,6 +807,41 @@ theorem getLsb_sshiftRight (x : BitVec w) (s i : Nat) :
807807
Nat.not_lt, decide_eq_true_eq]
808808
omega
809809

810+
/-- The msb after arithmetic shifting right equals the original msb. -/
811+
theorem sshiftRight_msb_eq_msb {n : Nat} {x : BitVec w} :
812+
(x.sshiftRight n).msb = x.msb := by
813+
rw [msb_eq_getLsb_last, getLsb_sshiftRight, msb_eq_getLsb_last]
814+
by_cases hw₀ : w = 0
815+
· simp [hw₀]
816+
· simp only [show ¬(w ≤ w - 1) by omega, decide_False, Bool.not_false, Bool.true_and,
817+
ite_eq_right_iff]
818+
intros h
819+
simp [show n = 0 by omega]
820+
821+
@[simp] theorem sshiftRight_zero {x : BitVec w} : x.sshiftRight 0 = x := by
822+
ext i
823+
simp
824+
825+
theorem sshiftRight_add {x : BitVec w} {m n : Nat} :
826+
x.sshiftRight (m + n) = (x.sshiftRight m).sshiftRight n := by
827+
ext i
828+
simp only [getLsb_sshiftRight, Nat.add_assoc]
829+
by_cases h₁ : w ≤ (i : Nat)
830+
· simp [h₁]
831+
· simp only [h₁, decide_False, Bool.not_false, Bool.true_and]
832+
by_cases h₂ : n + ↑i < w
833+
· simp [h₂]
834+
· simp only [h₂, ↓reduceIte]
835+
by_cases h₃ : m + (n + ↑i) < w
836+
· simp [h₃]
837+
omega
838+
· simp [h₃, sshiftRight_msb_eq_msb]
839+
840+
/-! ### sshiftRight reductions from BitVec to Nat -/
841+
842+
@[simp]
843+
theorem sshiftRight_eq' (x : BitVec w) : x.sshiftRight' y = x.sshiftRight y.toNat := rfl
844+
810845
/-! ### signExtend -/
811846

812847
/-- Equation theorem for `Int.sub` when both arguments are `Int.ofNat` -/

0 commit comments

Comments
 (0)