Skip to content

Commit 27cc0c8

Browse files
authored
feat: USize.reduceToNat (#6190)
This PR adds the builtin simproc `USize.reduceToNat` which reduces the `USize.toNat` operation on literals less than `UInt32.size` (i.e., `4294967296`).
1 parent c9ee66f commit 27cc0c8

File tree

2 files changed

+24
-3
lines changed

2 files changed

+24
-3
lines changed

src/Lean/Meta/Tactic/Simp/BuiltinSimprocs/UInt.lean

Lines changed: 17 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -84,8 +84,22 @@ declare_uint_simprocs UInt8
8484
declare_uint_simprocs UInt16
8585
declare_uint_simprocs UInt32
8686
declare_uint_simprocs UInt64
87+
8788
/-
88-
We disabled the simprocs for USize since the result of most operations depend on an opaque value: `System.Platform.numBits`.
89-
We could reduce some cases using the fact that this opaque value is `32` or `64`, but it is unclear whether it would be useful in practice.
89+
We do not use the normal simprocs for `USize` since the result of most operations depend on an opaque value: `System.Platform.numBits`.
90+
However, we do reduce natural literals using the fact this opaque value is at least `32`.
9091
-/
91-
-- declare_uint_simprocs USize
92+
namespace USize
93+
94+
def fromExpr (e : Expr) : SimpM (Option USize) := do
95+
let some (n, _) ← getOfNatValue? e ``USize | return none
96+
return USize.ofNat n
97+
98+
builtin_simproc [simp, seval] reduceToNat (USize.toNat _) := fun e => do
99+
let_expr USize.toNat e ← e | return .continue
100+
let some (n, _) ← getOfNatValue? e ``USize | return .continue
101+
unless n < UInt32.size do return .continue
102+
let e := toExpr n
103+
let p ← mkDecideProof (← mkLT e (mkNatLit UInt32.size))
104+
let p := mkApp2 (mkConst ``USize.toNat_ofNat_of_lt_32) e p
105+
return .done { expr := e, proof? := p }

tests/lean/run/simprocUInt.lean

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,7 @@
1+
variable (x : USize)
2+
3+
/- USize.toNat -/
4+
5+
#check_simp USize.toNat 4294967296 !~>
6+
#check_simp USize.toNat 4294967295 ~> 4294967295
7+
#check_simp USize.toNat (x &&& 1) ~> x.toNat % 2

0 commit comments

Comments
 (0)