Skip to content

Commit

Permalink
feat: add Float.toBits and Float.fromBits (#6094)
Browse files Browse the repository at this point in the history
This PR adds raw transmutation of floating-point numbers to and from
`UInt64`. Floats and UInts share the same endianness across all
supported platforms. The IEEE 754 standard precisely specifies the bit
layout of floats. Note that `Float.toBits` is distinct from
`Float.toUInt64`, which attempts to preserve the numeric value rather
than the bitwise value.

closes #6071
  • Loading branch information
leodemoura authored Nov 15, 2024
1 parent 691acde commit ecbaeff
Show file tree
Hide file tree
Showing 4 changed files with 49 additions and 0 deletions.
19 changes: 19 additions & 0 deletions src/Init/Data/Float.lean
Original file line number Diff line number Diff line change
Expand Up @@ -47,6 +47,25 @@ def Float.lt : Float → Float → Prop := fun a b =>
def Float.le : Float → Float → Prop := fun a b =>
floatSpec.le a.val b.val

/--
Raw transmutation from `UInt64`.
Floats and UInts have the same endianness on all supported platforms.
IEEE 754 very precisely specifies the bit layout of floats.
-/
@[extern "lean_float_from_bits"] opaque Float.fromBits : UInt64 → Float

/--
Raw transmutation to `UInt64`.
Floats and UInts have the same endianness on all supported platforms.
IEEE 754 very precisely specifies the bit layout of floats.
Note that this function is distinct from `Float.toUInt64`, which attempts
to preserve the numeric value, and not the bitwise value.
-/
@[extern "lean_float_to_bits"] opaque Float.toBits : Float → UInt64

instance : Add Float := ⟨Float.add⟩
instance : Sub Float := ⟨Float.sub⟩
instance : Mul Float := ⟨Float.mul⟩
Expand Down
2 changes: 2 additions & 0 deletions src/include/lean/lean.h
Original file line number Diff line number Diff line change
Expand Up @@ -2691,6 +2691,8 @@ static inline size_t lean_float_to_usize(double a) {
else
return (size_t) lean_float_to_uint32(a); // NOLINT
}
LEAN_EXPORT double lean_float_from_bits(uint64_t u);
LEAN_EXPORT uint64_t lean_float_to_bits(double d);
static inline double lean_float_add(double a, double b) { return a + b; }
static inline double lean_float_sub(double a, double b) { return a - b; }
static inline double lean_float_mul(double a, double b) { return a * b; }
Expand Down
15 changes: 15 additions & 0 deletions src/runtime/object.cpp
Original file line number Diff line number Diff line change
Expand Up @@ -1620,6 +1620,21 @@ extern "C" LEAN_EXPORT obj_res lean_float_frexp(double a) {
return r;
}

extern "C" LEAN_EXPORT double lean_float_from_bits(uint64_t u)
{
static_assert(sizeof(double) == sizeof(u), "`double` unexpected size.");
double ret;
std::memcpy(&ret, &u, sizeof(double));
return ret;
}

extern "C" LEAN_EXPORT uint64_t lean_float_to_bits(double d)
{
uint64_t ret;
std::memcpy(&ret, &d, sizeof(double));
return ret;
}

// =======================================
// Strings

Expand Down
13 changes: 13 additions & 0 deletions tests/lean/run/floatBits.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,13 @@
def d : Float := 1.245

/--
info: 4608285800708723180
-/
#guard_msgs in
#eval d.toBits

/--
info: true
-/
#guard_msgs in
#eval Float.fromBits d.toBits == d

0 comments on commit ecbaeff

Please sign in to comment.