Skip to content

Commit

Permalink
doc: update documentation and tests for toUIntX functions (#5497)
Browse files Browse the repository at this point in the history
Update documentation on functions to reflect actual behavior.
Add tests to ensure said behavior is as documented.

Closes #5483
  • Loading branch information
TomasPuverle authored Sep 29, 2024
1 parent cf3e7de commit 994cfa4
Show file tree
Hide file tree
Showing 3 changed files with 52 additions and 11 deletions.
36 changes: 25 additions & 11 deletions src/Init/Data/Float.lean
Original file line number Diff line number Diff line change
Expand Up @@ -72,21 +72,35 @@ instance floatDecLt (a b : Float) : Decidable (a < b) := Float.decLt a b
instance floatDecLe (a b : Float) : Decidable (a ≤ b) := Float.decLe a b

@[extern "lean_float_to_string"] opaque Float.toString : Float → String

/-- If the given float is positive, truncates the value to the nearest positive integer.
If negative or larger than the maximum value for UInt8, returns 0. -/
/-- If the given float is non-negative, truncates the value to the nearest non-negative integer.
If negative or NaN, returns 0.
If larger than the maximum value for UInt8 (including Inf), returns maximum value of UInt8
(i.e. UInt8.size - 1).
-/
@[extern "lean_float_to_uint8"] opaque Float.toUInt8 : Float → UInt8
/-- If the given float is positive, truncates the value to the nearest positive integer.
If negative or larger than the maximum value for UInt16, returns 0. -/
/-- If the given float is non-negative, truncates the value to the nearest non-negative integer.
If negative or NaN, returns 0.
If larger than the maximum value for UInt16 (including Inf), returns maximum value of UInt16
(i.e. UInt16.size - 1).
-/
@[extern "lean_float_to_uint16"] opaque Float.toUInt16 : Float → UInt16
/-- If the given float is positive, truncates the value to the nearest positive integer.
If negative or larger than the maximum value for UInt32, returns 0. -/
/-- If the given float is non-negative, truncates the value to the nearest non-negative integer.
If negative or NaN, returns 0.
If larger than the maximum value for UInt32 (including Inf), returns maximum value of UInt32
(i.e. UInt32.size - 1).
-/
@[extern "lean_float_to_uint32"] opaque Float.toUInt32 : Float → UInt32
/-- If the given float is positive, truncates the value to the nearest positive integer.
If negative or larger than the maximum value for UInt64, returns 0. -/
/-- If the given float is non-negative, truncates the value to the nearest non-negative integer.
If negative or NaN, returns 0.
If larger than the maximum value for UInt64 (including Inf), returns maximum value of UInt64
(i.e. UInt64.size - 1).
-/
@[extern "lean_float_to_uint64"] opaque Float.toUInt64 : Float → UInt64
/-- If the given float is positive, truncates the value to the nearest positive integer.
If negative or larger than the maximum value for USize, returns 0. -/
/-- If the given float is non-negative, truncates the value to the nearest non-negative integer.
If negative or NaN, returns 0.
If larger than the maximum value for USize (including Inf), returns maximum value of USize
(i.e. USize.size - 1; Note that this value is platform dependent).
-/
@[extern "lean_float_to_usize"] opaque Float.toUSize : Float → USize

@[extern "lean_float_isnan"] opaque Float.isNaN : Float → Bool
Expand Down
17 changes: 17 additions & 0 deletions tests/compiler/float.lean
Original file line number Diff line number Diff line change
Expand Up @@ -19,12 +19,29 @@ def tst1 : IO Unit := do
IO.println (0 / 0 : Float).toUInt16
IO.println (0 / 0 : Float).toUInt32
IO.println (0 / 0 : Float).toUInt64
IO.println (0 / 0 : Float).toUSize

IO.println (-1 : Float).toUInt8
IO.println (256 : Float).toUInt8
IO.println (1 / 0 : Float).toUInt8

IO.println (-1 : Float).toUInt16
IO.println (2^16 : Float).toUInt16
IO.println (1 / 0 : Float).toUInt16

IO.println (-1 : Float).toUInt32
IO.println (2^32 : Float).toUInt32
IO.println (1 / 0 : Float).toUInt32

IO.println (-1 : Float).toUInt64
IO.println (2^64 : Float).toUInt64
IO.println (1 / 0 : Float).toUInt64

let platBits := USize.size.log2.toFloat
IO.println (-1 : Float).toUSize
IO.println ((2^platBits).toUSize.toNat == (USize.size - 1))
IO.println ((1 / 0 : Float).toUSize.toNat == (USize.size - 1))

IO.println (let x : Float := 1.4; (x, x.isNaN, x.isInf, x.isFinite, x.frExp))
IO.println (let x : Float := 0 / 0; (x, x.isNaN, x.isInf, x.isFinite, x.frExp))
IO.println (let x : Float := -0 / 0; (x, x.isNaN, x.isInf, x.isFinite, x.frExp))
Expand Down
10 changes: 10 additions & 0 deletions tests/compiler/float.lean.expected.out
Original file line number Diff line number Diff line change
Expand Up @@ -18,11 +18,21 @@ true
0
0
0
0
255
255
0
65535
65535
0
4294967295
4294967295
0
18446744073709551615
18446744073709551615
0
true
true
(1.400000, (false, (false, (true, (0.700000, 1)))))
(NaN, (true, (false, (false, (NaN, 0)))))
(NaN, (true, (false, (false, (NaN, 0)))))
Expand Down

0 comments on commit 994cfa4

Please sign in to comment.