Skip to content

Commit c6863bd

Browse files
bolluJovanGerb
authored andcommitted
feat: BitVec lemmas for smtUDiv, smtSDiv when denominator is zero (leanprover#5616)
This PR is a follow-up to leanprover#5609, where we add lemmas characterizing `smtUDiv` and `smtSDiv`'s behavior when the denominator is zero. We build some `slt` theory, connecting it to `msb` for a clean proof. I chose not to characterize `slt` in terms of `msb` a `simp` lemma, since I anticipate use cases where we want to keep the arithmetic interpretation of `slt`.
1 parent 179ba7b commit c6863bd

File tree

1 file changed

+32
-0
lines changed

1 file changed

+32
-0
lines changed

src/Init/Data/BitVec/Lemmas.lean

Lines changed: 32 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -512,6 +512,31 @@ theorem eq_zero_or_eq_one (a : BitVec 1) : a = 0#1 ∨ a = 1#1 := by
512512
subst h
513513
simp
514514

515+
@[simp]
516+
theorem toInt_zero {w : Nat} : (0#w).toInt = 0 := by
517+
simp [BitVec.toInt, show 0 < 2^w by exact Nat.two_pow_pos w]
518+
519+
/-! ### slt -/
520+
521+
/--
522+
A bitvector, when interpreted as an integer, is less than zero iff
523+
its most significant bit is true.
524+
-/
525+
theorem slt_zero_iff_msb_cond (x : BitVec w) : x.slt 0#w ↔ x.msb = true := by
526+
have := toInt_eq_msb_cond x
527+
constructor
528+
· intros h
529+
apply Classical.byContradiction
530+
intros hmsb
531+
simp only [Bool.not_eq_true] at hmsb
532+
simp only [hmsb, Bool.false_eq_true, ↓reduceIte] at this
533+
simp only [BitVec.slt, toInt_zero, decide_eq_true_eq] at h
534+
omega /- Can't have `x.toInt` which is equal to `x.toNat` be strictly less than zero -/
535+
· intros h
536+
simp only [h, ↓reduceIte] at this
537+
simp [BitVec.slt, this]
538+
omega
539+
515540
/-! ### setWidth, zeroExtend and truncate -/
516541

517542
@[simp]
@@ -2376,6 +2401,9 @@ theorem umod_eq_and {x y : BitVec 1} : x % y = x &&& (~~~y) := by
23762401
theorem smtUDiv_eq (x y : BitVec w) : smtUDiv x y = if y = 0#w then allOnes w else x / y := by
23772402
simp [smtUDiv]
23782403

2404+
@[simp]
2405+
theorem smtUDiv_zero {x : BitVec n} : x.smtUDiv 0#n = allOnes n := rfl
2406+
23792407
/-! ### sdiv -/
23802408

23812409
/-- Equation theorem for `sdiv` in terms of `udiv`. -/
@@ -2443,6 +2471,10 @@ theorem smtSDiv_eq (x y : BitVec w) : smtSDiv x y =
24432471
rw [BitVec.smtSDiv]
24442472
rcases x.msb <;> rcases y.msb <;> simp
24452473

2474+
@[simp]
2475+
theorem smtSDiv_zero {x : BitVec n} : x.smtSDiv 0#n = if x.slt 0#n then 1#n else (allOnes n) := by
2476+
rcases hx : x.msb <;> simp [smtSDiv, slt_zero_iff_msb_cond x, hx, ← negOne_eq_allOnes]
2477+
24462478
/-! ### srem -/
24472479

24482480
theorem srem_eq (x y : BitVec w) : srem x y =

0 commit comments

Comments
 (0)