Skip to content

Commit 4c20035

Browse files
alexkeizerkim-em
authored andcommitted
feat: BitVec.{toFin, toInt, msb}_umod (leanprover#6404)
This PR adds a `toFin` and `msb` lemma for unsigned bitvector modulus. Similar to leanprover#6402, we don't provide a general `toInt_umod` lemmas, but instead choose to provide more specialized rewrites, with extra side-conditions. --------- Co-authored-by: Kim Morrison <[email protected]>
1 parent 9c9cab8 commit 4c20035

File tree

2 files changed

+65
-0
lines changed

2 files changed

+65
-0
lines changed

src/Init/Data/BitVec/Lemmas.lean

+60
Original file line numberDiff line numberDiff line change
@@ -11,6 +11,7 @@ import Init.Data.Fin.Lemmas
1111
import Init.Data.Nat.Lemmas
1212
import Init.Data.Nat.Div.Lemmas
1313
import Init.Data.Nat.Mod
14+
import Init.Data.Nat.Div.Lemmas
1415
import Init.Data.Int.Bitwise.Lemmas
1516
import Init.Data.Int.Pow
1617

@@ -99,6 +100,12 @@ theorem ofFin_eq_ofNat : @BitVec.ofFin w (Fin.mk x lt) = BitVec.ofNat w x := by
99100
theorem eq_of_toNat_eq {n} : ∀ {x y : BitVec n}, x.toNat = y.toNat → x = y
100101
| ⟨_, _⟩, ⟨_, _⟩, rfl => rfl
101102

103+
/-- Prove nonequality of bitvectors in terms of nat operations. -/
104+
theorem toNat_ne_iff_ne {n} {x y : BitVec n} : x.toNat ≠ y.toNat ↔ x ≠ y := by
105+
constructor
106+
· rintro h rfl; apply h rfl
107+
· intro h h_eq; apply h <| eq_of_toNat_eq h_eq
108+
102109
@[simp] theorem val_toFin (x : BitVec w) : x.toFin.val = x.toNat := rfl
103110

104111
@[bv_toNat] theorem toNat_eq {x y : BitVec n} : x = y ↔ x.toNat = y.toNat :=
@@ -2693,6 +2700,10 @@ theorem umod_def {x y : BitVec n} :
26932700
theorem toNat_umod {x y : BitVec n} :
26942701
(x % y).toNat = x.toNat % y.toNat := rfl
26952702

2703+
@[simp]
2704+
theorem toFin_umod {x y : BitVec w} :
2705+
(x % y).toFin = x.toFin % y.toFin := rfl
2706+
26962707
@[simp]
26972708
theorem umod_zero {x : BitVec n} : x % 0#n = x := by
26982709
simp [umod_def]
@@ -2720,6 +2731,55 @@ theorem umod_eq_and {x y : BitVec 1} : x % y = x &&& (~~~y) := by
27202731
rcases hy with rfl | rfl <;>
27212732
rfl
27222733

2734+
theorem umod_eq_of_lt {x y : BitVec w} (h : x < y) :
2735+
x % y = x := by
2736+
apply eq_of_toNat_eq
2737+
simp [Nat.mod_eq_of_lt h]
2738+
2739+
@[simp]
2740+
theorem msb_umod {x y : BitVec w} :
2741+
(x % y).msb = (x.msb && (x < y || y == 0#w)) := by
2742+
rw [msb_eq_decide, toNat_umod]
2743+
cases msb_x : x.msb
2744+
· suffices x.toNat % y.toNat < 2 ^ (w - 1) by simpa
2745+
calc
2746+
x.toNat % y.toNat ≤ x.toNat := by apply Nat.mod_le
2747+
_ < 2 ^ (w - 1) := by simpa [msb_eq_decide] using msb_x
2748+
. by_cases hy : y = 0
2749+
· simp_all [msb_eq_decide]
2750+
· suffices 2 ^ (w - 1) ≤ x.toNat % y.toNat ↔ x < y by simp_all
2751+
by_cases x_lt_y : x < y
2752+
. simp_all [Nat.mod_eq_of_lt x_lt_y, msb_eq_decide]
2753+
· suffices x.toNat % y.toNat < 2 ^ (w - 1) by
2754+
simpa [x_lt_y]
2755+
have y_le_x : y.toNat ≤ x.toNat := by
2756+
simpa using x_lt_y
2757+
replace hy : y.toNat ≠ 0 :=
2758+
toNat_ne_iff_ne.mpr hy
2759+
by_cases msb_y : y.toNat < 2 ^ (w - 1)
2760+
· have : x.toNat % y.toNat < y.toNat := Nat.mod_lt _ (by omega)
2761+
omega
2762+
· rcases w with _|w
2763+
· contradiction
2764+
simp only [Nat.add_one_sub_one]
2765+
replace msb_y : 2 ^ w ≤ y.toNat := by
2766+
simpa using msb_y
2767+
have : y.toNat ≤ y.toNat * (x.toNat / y.toNat) := by
2768+
apply Nat.le_mul_of_pos_right
2769+
apply Nat.div_pos y_le_x
2770+
omega
2771+
have : x.toNat % y.toNat ≤ x.toNat - y.toNat := by
2772+
rw [Nat.mod_eq_sub]; omega
2773+
omega
2774+
2775+
theorem toInt_umod {x y : BitVec w} :
2776+
(x % y).toInt = (x.toNat % y.toNat : Int).bmod (2 ^ w) := by
2777+
simp [toInt_eq_toNat_bmod]
2778+
2779+
theorem toInt_umod_of_msb {x y : BitVec w} (h : x.msb = false) :
2780+
(x % y).toInt = x.toInt % y.toNat := by
2781+
simp [toInt_eq_msb_cond, h]
2782+
27232783
/-! ### smtUDiv -/
27242784

27252785
theorem smtUDiv_eq (x y : BitVec w) : smtUDiv x y = if y = 0#w then allOnes w else x / y := by

src/Init/Data/Nat/Div/Lemmas.lean

+5
Original file line numberDiff line numberDiff line change
@@ -49,6 +49,11 @@ theorem lt_div_mul_self (h : 0 < k) (w : k ≤ x) : x - k < x / k * k := by
4949
have : x % k < k := mod_lt x h
5050
omega
5151

52+
theorem div_pos (hba : b ≤ a) (hb : 0 < b) : 0 < a / b := by
53+
cases b
54+
· contradiction
55+
· simp [Nat.pos_iff_ne_zero, div_eq_zero_iff_lt, hba]
56+
5257
theorem div_le_div_left (hcb : c ≤ b) (hc : 0 < c) : a / b ≤ a / c :=
5358
(Nat.le_div_iff_mul_le hc).2 <|
5459
Nat.le_trans (Nat.mul_le_mul_left _ hcb) (Nat.div_mul_le_self a b)

0 commit comments

Comments
 (0)