Skip to content

Commit

Permalink
feat: umod_eq_of_lt
Browse files Browse the repository at this point in the history
  • Loading branch information
alexkeizer committed Dec 16, 2024
1 parent 03373e7 commit d16ff31
Showing 1 changed file with 11 additions and 0 deletions.
11 changes: 11 additions & 0 deletions src/Init/Data/BitVec/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -99,6 +99,12 @@ theorem ofFin_eq_ofNat : @BitVec.ofFin w (Fin.mk x lt) = BitVec.ofNat w x := by
theorem eq_of_toNat_eq {n} : ∀ {x y : BitVec n}, x.toNat = y.toNat → x = y
| ⟨_, _⟩, ⟨_, _⟩, rfl => rfl

/-- Prove nonequality of bitvectors in terms of nat operations. -/
theorem neq_iff_toNat_neq {n} {x y : BitVec n} : x.toNat ≠ y.toNat ↔ x ≠ y := by
constructor
· rintro h rfl; apply h rfl
· intro h h_eq; apply h <| eq_of_toNat_eq h_eq

@[simp] theorem val_toFin (x : BitVec w) : x.toFin.val = x.toNat := rfl

@[bv_toNat] theorem toNat_eq {x y : BitVec n} : x = y ↔ x.toNat = y.toNat :=
Expand Down Expand Up @@ -2606,6 +2612,11 @@ theorem umod_eq_and {x y : BitVec 1} : x % y = x &&& (~~~y) := by
rcases hy with rfl | rfl <;>
rfl

theorem umod_eq_of_lt {x y : BitVec w} (h : x < y) :
x % y = x := by
apply eq_of_toNat_eq
simp [Nat.mod_eq_of_lt h]

@[simp]
theorem msb_umod {x y : BitVec w} :
(x % y).msb = (x.msb && (x < y || y == 0#w)) := by
Expand Down

0 comments on commit d16ff31

Please sign in to comment.