Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

feat: add the int_toBitVec simp set #6652

Merged
merged 3 commits into from
Jan 15, 2025
Merged
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
32 changes: 16 additions & 16 deletions src/Init/Data/UInt/Bitwise.lean
Original file line number Diff line number Diff line change
Expand Up @@ -13,17 +13,17 @@ macro "declare_bitwise_uint_theorems" typeName:ident bits:term:arg : command =>
`(
namespace $typeName

@[simp] protected theorem toBitVec_add {a b : $typeName} : (a + b).toBitVec = a.toBitVec + b.toBitVec := rfl
@[simp] protected theorem toBitVec_sub {a b : $typeName} : (a - b).toBitVec = a.toBitVec - b.toBitVec := rfl
@[simp] protected theorem toBitVec_mul {a b : $typeName} : (a * b).toBitVec = a.toBitVec * b.toBitVec := rfl
@[simp] protected theorem toBitVec_div {a b : $typeName} : (a / b).toBitVec = a.toBitVec / b.toBitVec := rfl
@[simp] protected theorem toBitVec_mod {a b : $typeName} : (a % b).toBitVec = a.toBitVec % b.toBitVec := rfl
@[simp] protected theorem toBitVec_not {a : $typeName} : (~~~a).toBitVec = ~~~a.toBitVec := rfl
@[simp] protected theorem toBitVec_and (a b : $typeName) : (a &&& b).toBitVec = a.toBitVec &&& b.toBitVec := rfl
@[simp] protected theorem toBitVec_or (a b : $typeName) : (a ||| b).toBitVec = a.toBitVec ||| b.toBitVec := rfl
@[simp] protected theorem toBitVec_xor (a b : $typeName) : (a ^^^ b).toBitVec = a.toBitVec ^^^ b.toBitVec := rfl
@[simp] protected theorem toBitVec_shiftLeft (a b : $typeName) : (a <<< b).toBitVec = a.toBitVec <<< (b.toBitVec % $bits) := rfl
@[simp] protected theorem toBitVec_shiftRight (a b : $typeName) : (a >>> b).toBitVec = a.toBitVec >>> (b.toBitVec % $bits) := rfl
@[simp, int_toBitVec] protected theorem toBitVec_add {a b : $typeName} : (a + b).toBitVec = a.toBitVec + b.toBitVec := rfl
@[simp, int_toBitVec] protected theorem toBitVec_sub {a b : $typeName} : (a - b).toBitVec = a.toBitVec - b.toBitVec := rfl
@[simp, int_toBitVec] protected theorem toBitVec_mul {a b : $typeName} : (a * b).toBitVec = a.toBitVec * b.toBitVec := rfl
@[simp, int_toBitVec] protected theorem toBitVec_div {a b : $typeName} : (a / b).toBitVec = a.toBitVec / b.toBitVec := rfl
@[simp, int_toBitVec] protected theorem toBitVec_mod {a b : $typeName} : (a % b).toBitVec = a.toBitVec % b.toBitVec := rfl
@[simp, int_toBitVec] protected theorem toBitVec_not {a : $typeName} : (~~~a).toBitVec = ~~~a.toBitVec := rfl
@[simp, int_toBitVec] protected theorem toBitVec_and (a b : $typeName) : (a &&& b).toBitVec = a.toBitVec &&& b.toBitVec := rfl
@[simp, int_toBitVec] protected theorem toBitVec_or (a b : $typeName) : (a ||| b).toBitVec = a.toBitVec ||| b.toBitVec := rfl
@[simp, int_toBitVec] protected theorem toBitVec_xor (a b : $typeName) : (a ^^^ b).toBitVec = a.toBitVec ^^^ b.toBitVec := rfl
@[simp, int_toBitVec] protected theorem toBitVec_shiftLeft (a b : $typeName) : (a <<< b).toBitVec = a.toBitVec <<< (b.toBitVec % $bits) := rfl
@[simp, int_toBitVec] protected theorem toBitVec_shiftRight (a b : $typeName) : (a >>> b).toBitVec = a.toBitVec >>> (b.toBitVec % $bits) := rfl

@[simp] protected theorem toNat_and (a b : $typeName) : (a &&& b).toNat = a.toNat &&& b.toNat := by simp [toNat]
@[simp] protected theorem toNat_or (a b : $typeName) : (a ||| b).toNat = a.toNat ||| b.toNat := by simp [toNat]
Expand All @@ -44,27 +44,27 @@ declare_bitwise_uint_theorems UInt32 32
declare_bitwise_uint_theorems UInt64 64
declare_bitwise_uint_theorems USize System.Platform.numBits

@[simp]
@[simp, int_toBitVec]
theorem Bool.toBitVec_toUInt8 {b : Bool} :
b.toUInt8.toBitVec = (BitVec.ofBool b).setWidth 8 := by
cases b <;> simp [toUInt8]

@[simp]
@[simp, int_toBitVec]
theorem Bool.toBitVec_toUInt16 {b : Bool} :
b.toUInt16.toBitVec = (BitVec.ofBool b).setWidth 16 := by
cases b <;> simp [toUInt16]

@[simp]
@[simp, int_toBitVec]
theorem Bool.toBitVec_toUInt32 {b : Bool} :
b.toUInt32.toBitVec = (BitVec.ofBool b).setWidth 32 := by
cases b <;> simp [toUInt32]

@[simp]
@[simp, int_toBitVec]
theorem Bool.toBitVec_toUInt64 {b : Bool} :
b.toUInt64.toBitVec = (BitVec.ofBool b).setWidth 64 := by
cases b <;> simp [toUInt64]

@[simp]
@[simp, int_toBitVec]
theorem Bool.toBitVec_toUSize {b : Bool} :
b.toUSize.toBitVec = (BitVec.ofBool b).setWidth System.Platform.numBits := by
cases b
Expand Down
20 changes: 17 additions & 3 deletions src/Init/Data/UInt/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -41,9 +41,9 @@ macro "declare_uint_theorems" typeName:ident bits:term:arg : command => do
theorem toNat_ofNat_of_lt {n : Nat} (h : n < size) : (ofNat n).toNat = n := by
rw [toNat, toBitVec_eq_of_lt h]

theorem le_def {a b : $typeName} : a ≤ b ↔ a.toBitVec ≤ b.toBitVec := .rfl
@[int_toBitVec] theorem le_def {a b : $typeName} : a ≤ b ↔ a.toBitVec ≤ b.toBitVec := .rfl

theorem lt_def {a b : $typeName} : a < b ↔ a.toBitVec < b.toBitVec := .rfl
@[int_toBitVec] theorem lt_def {a b : $typeName} : a < b ↔ a.toBitVec < b.toBitVec := .rfl

theorem le_iff_toNat_le {a b : $typeName} : a ≤ b ↔ a.toNat ≤ b.toNat := .rfl

Expand Down Expand Up @@ -74,6 +74,11 @@ macro "declare_uint_theorems" typeName:ident bits:term:arg : command => do
protected theorem toBitVec_inj {a b : $typeName} : a.toBitVec = b.toBitVec ↔ a = b :=
Iff.intro eq_of_toBitVec_eq toBitVec_eq_of_eq

open $typeName (eq_of_toBitVec_eq toBitVec_eq_of_eq) in
@[int_toBitVec]
protected theorem eq_iff_toBitVec_eq {a b : $typeName} : a = b ↔ a.toBitVec = b.toBitVec :=
Iff.intro toBitVec_eq_of_eq eq_of_toBitVec_eq

open $typeName (eq_of_toBitVec_eq) in
protected theorem eq_of_val_eq {a b : $typeName} (h : a.val = b.val) : a = b := by
rcases a with ⟨⟨_⟩⟩; rcases b with ⟨⟨_⟩⟩; simp_all [val]
Expand All @@ -82,10 +87,19 @@ macro "declare_uint_theorems" typeName:ident bits:term:arg : command => do
protected theorem val_inj {a b : $typeName} : a.val = b.val ↔ a = b :=
Iff.intro eq_of_val_eq (congrArg val)

open $typeName (eq_of_toBitVec_eq) in
protected theorem toBitVec_ne_of_ne {a b : $typeName} (h : a ≠ b) : a.toBitVec ≠ b.toBitVec :=
fun h' => h (eq_of_toBitVec_eq h')

open $typeName (toBitVec_eq_of_eq) in
protected theorem ne_of_toBitVec_ne {a b : $typeName} (h : a.toBitVec ≠ b.toBitVec) : a ≠ b :=
fun h' => absurd (toBitVec_eq_of_eq h') h

open $typeName (ne_of_toBitVec_ne toBitVec_ne_of_ne) in
@[int_toBitVec]
protected theorem ne_iff_toBitVec_ne {a b : $typeName} : a ≠ b ↔ a.toBitVec ≠ b.toBitVec :=
Iff.intro toBitVec_ne_of_ne ne_of_toBitVec_ne

open $typeName (ne_of_toBitVec_ne) in
protected theorem ne_of_lt {a b : $typeName} (h : a < b) : a ≠ b := by
apply ne_of_toBitVec_ne
Expand Down Expand Up @@ -159,7 +173,7 @@ macro "declare_uint_theorems" typeName:ident bits:term:arg : command => do
@[simp]
theorem val_ofNat (n : Nat) : val (no_index (OfNat.ofNat n)) = OfNat.ofNat n := rfl

@[simp]
@[simp, int_toBitVec]
theorem toBitVec_ofNat (n : Nat) : toBitVec (no_index (OfNat.ofNat n)) = BitVec.ofNat _ n := rfl

@[simp]
Expand Down
3 changes: 3 additions & 0 deletions src/Lean/Elab/Tactic/BVDecide/Frontend/Attr.lean
Original file line number Diff line number Diff line change
Expand Up @@ -38,6 +38,9 @@ declare_config_elab elabBVDecideConfig Lean.Elab.Tactic.BVDecide.Frontend.BVDeci
builtin_initialize bvNormalizeExt : Meta.SimpExtension ←
Meta.registerSimpAttr `bv_normalize "simp theorems used by bv_normalize"

builtin_initialize intToBitVecExt : Meta.SimpExtension ←
Meta.registerSimpAttr `int_toBitVec "simp theorems used to convert UIntX/IntX statements into BitVec ones"

/-- Builtin `bv_normalize` simprocs. -/
builtin_initialize builtinBVNormalizeSimprocsRef : IO.Ref Meta.Simp.Simprocs ← IO.mkRef {}

Expand Down
3 changes: 2 additions & 1 deletion stage0/src/runtime/CMakeLists.txt

Some generated files are not rendered by default. Learn more about how customized files appear on GitHub.

2 changes: 2 additions & 0 deletions stage0/src/runtime/init_module.cpp

Some generated files are not rendered by default. Learn more about how customized files appear on GitHub.

19 changes: 17 additions & 2 deletions stage0/src/runtime/libuv.cpp

Some generated files are not rendered by default. Learn more about how customized files appear on GitHub.

22 changes: 21 additions & 1 deletion stage0/src/runtime/libuv.h

Some generated files are not rendered by default. Learn more about how customized files appear on GitHub.

5 changes: 5 additions & 0 deletions stage0/src/runtime/object.h

Some generated files are not rendered by default. Learn more about how customized files appear on GitHub.

143 changes: 143 additions & 0 deletions stage0/src/runtime/uv/event_loop.cpp

Some generated files are not rendered by default. Learn more about how customized files appear on GitHub.

Loading
Loading