@@ -1662,6 +1662,88 @@ theorem toInt_sdiv (a b : BitVec w) : (a.sdiv b).toInt = (a.toInt.tdiv b.toInt).
16621662 · rw [← toInt_bmod_cancel]
16631663 rw [BitVec.toInt_sdiv_of_ne_or_ne _ _ (by simpa only [Decidable.not_and_iff_not_or_not] using h)]
16641664
1665+ private theorem neg_udiv_eq_intMin_iff_eq_intMin_eq_one_of_msb_eq_true
1666+ {x y : BitVec w} (hx : x.msb = true) (hy : y.msb = false) :
1667+ -x / y = intMin w ↔ (x = intMin w ∧ y = 1 #w) := by
1668+ constructor
1669+ · intros h
1670+ rcases w with _ | w; decide +revert
1671+ have : (-x / y).msb = true := by simp [h, msb_intMin]
1672+ rw [msb_udiv] at this
1673+ simp only [bool_to_prop] at this
1674+ obtain ⟨hx, hy⟩ := this
1675+ simp only [beq_iff_eq] at hy
1676+ subst hy
1677+ simp only [udiv_one, zero_lt_succ, neg_eq_intMin] at h
1678+ simp [h]
1679+ · rintro ⟨hx, hy⟩
1680+ subst hx hy
1681+ simp
1682+
1683+ /--
1684+ the most significant bit of the signed division `x.sdiv y` can be computed
1685+ by the following cases:
1686+ (1) x nonneg, y nonneg: never neg.
1687+ (2) x nonneg, y neg: neg when result nonzero.
1688+ We know that y is nonzero since it is negative, so we only check `|x| ≥ |y|`.
1689+ (3) x neg, y nonneg: neg when result nonzero.
1690+ We check that `y ≠ 0` and `|x| ≥ |y|`.
1691+ (4) x neg, y neg: neg when `x = intMin, `y = -1`, since `intMin / -1 = intMin`.
1692+
1693+ The proof strategy is to perform a case analysis on the sign of `x` and `y`,
1694+ followed by unfolding the `sdiv` into `udiv`.
1695+ -/
1696+ theorem msb_sdiv_eq_decide {x y : BitVec w} :
1697+ (x.sdiv y).msb = (decide (0 < w) &&
1698+ (!x.msb && y.msb && decide (-y ≤ x)) ||
1699+ (x.msb && !y.msb && decide (y ≤ -x) && !decide (y = 0 #w)) ||
1700+ (x.msb && y.msb && decide (x = intMin w) && decide (y = -1 #w)))
1701+ := by
1702+ rcases w; decide +revert
1703+ case succ w =>
1704+ simp only [decide_true, ne_eq, decide_and, decide_not, Bool.true_and,
1705+ sdiv_eq, udiv_eq]
1706+ rcases hxmsb : x.msb <;> rcases hymsb : y.msb
1707+ · simp [hxmsb, hymsb, msb_udiv_eq_false_of, Bool.not_false, Bool.and_false, Bool.false_and,
1708+ Bool.and_true, Bool.or_self, Bool.and_self]
1709+ · simp only [hxmsb, hymsb, msb_neg, msb_udiv_eq_false_of, bne_false, Bool.not_false,
1710+ Bool.and_self, ne_zero_of_msb_true, decide_false, Bool.and_true, Bool.true_and, Bool.not_true,
1711+ Bool.false_and, Bool.or_false, bool_to_prop]
1712+ have : x / -y ≠ intMin (w + 1 ) := by
1713+ intros h
1714+ have : (x / -y).msb = (intMin (w + 1 )).msb := by simp only [h]
1715+ simp only [msb_udiv, msb_intMin, show 0 < w + 1 by omega, decide_true, and_eq_true, beq_iff_eq] at this
1716+ obtain ⟨hcontra, _⟩ := this
1717+ simp only [hcontra, true_eq_false] at hxmsb
1718+ simp [this , hymsb, udiv_ne_zero_iff_ne_zero_and_le]
1719+ · simp only [hxmsb, hymsb, Bool.not_true, Bool.and_self, Bool.false_and, Bool.not_false,
1720+ Bool.true_and, Bool.false_or, Bool.and_false, Bool.or_false]
1721+ by_cases hx₁ : x = 0 #(w + 1 )
1722+ · simp [hx₁, neg_zero, zero_udiv, msb_zero, le_zero_iff, Bool.and_not_self]
1723+ · by_cases hy₁ : y = 0 #(w + 1 )
1724+ · simp [hy₁, udiv_zero, neg_zero, msb_zero, decide_true, Bool.not_true, Bool.and_false]
1725+ · simp only [hy₁, decide_false, Bool.not_false, Bool.and_true]
1726+ by_cases hxy₁ : (- x / y) = 0 #(w + 1 )
1727+ · simp only [hxy₁, neg_zero, msb_zero, false_eq_decide_iff, BitVec.not_le,
1728+ decide_eq_true_eq, BitVec.not_le]
1729+ simp only [udiv_eq_zero_iff_eq_zero_or_lt, hy₁, _root_.false_or] at hxy₁
1730+ bv_omega
1731+ · simp only [udiv_eq_zero_iff_eq_zero_or_lt, _root_.not_or, BitVec.not_lt,
1732+ hy₁, not_false_eq_true, _root_.true_and] at hxy₁
1733+ simp only [hxy₁, decide_true, msb_neg, bne_iff_ne, ne_eq,
1734+ bool_to_prop,
1735+ bne_iff_ne, ne_eq, udiv_eq_zero_iff_eq_zero_or_lt, hy₁, _root_.false_or,
1736+ BitVec.not_lt, hxy₁, _root_.true_and, decide_not, not_eq_eq_eq_not, not_eq_not,
1737+ msb_udiv, msb_neg]
1738+ simp only [hx₁, not_false_eq_true, _root_.true_and, decide_not, hxmsb, not_eq_eq_eq_not,
1739+ Bool.not_true, decide_eq_false_iff_not, Decidable.not_not, beq_iff_eq]
1740+ rw [neg_udiv_eq_intMin_iff_eq_intMin_eq_one_of_msb_eq_true hxmsb hymsb]
1741+ · simp only [msb_udiv, msb_neg, hxmsb, bne_true, Bool.not_and, Bool.not_true, Bool.and_true,
1742+ Bool.false_and, Bool.and_false, hymsb, ne_zero_of_msb_true, decide_false, Bool.not_false,
1743+ Bool.or_self, Bool.and_self, Bool.true_and, Bool.false_or]
1744+ simp only [bool_to_prop]
1745+ simp [BitVec.ne_zero_of_msb_true (x := x) hxmsb, neg_eq_iff_eq_neg]
1746+
16651747theorem msb_umod_eq_false_of_left {x : BitVec w} (hx : x.msb = false) (y : BitVec w) : (x % y).msb = false := by
16661748 rw [msb_eq_false_iff_two_mul_lt] at hx ⊢
16671749 rw [toNat_umod]
0 commit comments