@@ -429,6 +429,67 @@ theorem shiftLeft_eq_shiftLeftRec (x : BitVec w₁) (y : BitVec w₂) :
429
429
· simp [of_length_zero]
430
430
· simp [shiftLeftRec_eq]
431
431
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
+
432
493
/- ### Logical shift right (ushiftRight) recurrence for bitblasting -/
433
494
434
495
/--
0 commit comments