diff options
author | Rutger Broekhoff | 2025-07-07 21:52:08 +0200 |
---|---|---|
committer | Rutger Broekhoff | 2025-07-07 21:52:08 +0200 |
commit | ba61dfd69504ec6263a9dee9931d93adeb6f3142 (patch) | |
tree | d6c9b78e50eeab24e0c1c09ab45909a6ae3fd5db /theories/nix/floats.v | |
download | verified-dyn-lang-interp-ba61dfd69504ec6263a9dee9931d93adeb6f3142.tar.gz verified-dyn-lang-interp-ba61dfd69504ec6263a9dee9931d93adeb6f3142.zip |
Initialize repository
Diffstat (limited to 'theories/nix/floats.v')
-rw-r--r-- | theories/nix/floats.v | 85 |
1 files changed, 85 insertions, 0 deletions
diff --git a/theories/nix/floats.v b/theories/nix/floats.v new file mode 100644 index 0000000..246e0c3 --- /dev/null +++ b/theories/nix/floats.v | |||
@@ -0,0 +1,85 @@ | |||
1 | From stdpp Require Import prelude ssreflect. | ||
2 | From Flocq.IEEE754 Require Import | ||
3 | Binary BinarySingleNaN (mode_NE, mode_DN, mode_UP) Bits. | ||
4 | From stdpp Require Import options. | ||
5 | |||
6 | Global Arguments B754_zero {_ _}. | ||
7 | Global Arguments B754_infinity {_ _}. | ||
8 | Global Arguments B754_nan {_ _}. | ||
9 | Global Arguments B754_finite {_ _}. | ||
10 | |||
11 | (** The bit representation of floats is not observable in Nix, and it appears | ||
12 | that only negative NaNs are ever produced. So we setup the Flocq floats in | ||
13 | the way that it always produces [-NaN], i.e., [indef_nan]. *) | ||
14 | Definition float := binary64. | ||
15 | |||
16 | Variant round_mode := Floor | Ceil | NearestEven. | ||
17 | Global Instance round_mode_eq_dec : EqDecision round_mode. | ||
18 | Proof. solve_decision. Defined. | ||
19 | |||
20 | Module Float. | ||
21 | Definition prec : Z := 53. | ||
22 | Definition emax : Z := 1024. | ||
23 | |||
24 | Lemma Hprec : (0 < 53)%Z. | ||
25 | Proof. done. Qed. | ||
26 | Lemma Hprec_emax : (53 < 1024)%Z. | ||
27 | Proof. done. Qed. | ||
28 | |||
29 | Global Instance inhabited : Inhabited float := populate (B754_zero false). | ||
30 | |||
31 | Global Instance eq_dec : EqDecision float. | ||
32 | Proof. | ||
33 | refine (λ f1 f2, | ||
34 | match f1, f2 with | ||
35 | | B754_zero s1, B754_zero s2 => cast_if (decide (s1 = s2)) | ||
36 | | B754_infinity s1, B754_infinity s2 => cast_if (decide (s1 = s2)) | ||
37 | | B754_nan s1 pl1 _, B754_nan s2 pl2 _ => | ||
38 | cast_if_and (decide (s1 = s2)) (decide (pl1 = pl2)) | ||
39 | | B754_finite s1 m1 e1 _, B754_finite s2 m2 e2 _ => | ||
40 | cast_if_and3 (decide (s1 = s2)) (decide (m1 = m2)) (decide (e1 = e2)) | ||
41 | | _, _ => right _ | ||
42 | end); abstract naive_solver (f_equal; apply (proof_irrel _)). | ||
43 | Defined. | ||
44 | |||
45 | Definition of_Z (x : Z) : float := | ||
46 | binary_normalize prec emax (refl_equal _) (refl_equal _) mode_NE x 0 false. | ||
47 | |||
48 | Definition to_Z (f : float) : option Z := | ||
49 | match f with | ||
50 | | B754_zero _ => Some 0 | ||
51 | | B754_finite s m e _ => Some (Zaux.cond_Zopp s (Zpos m) ≪ e) | ||
52 | | _ => None | ||
53 | end%Z. | ||
54 | |||
55 | (** QNaN Floating-Point Indefinite; see Table 4-3. Floating-Point Number and | ||
56 | NaN Encodings. *) | ||
57 | Definition indef_nan : { f | is_nan prec emax f = true } := | ||
58 | @B754_nan prec emax true (2^51) (refl_equal _) ↾ eq_refl _. | ||
59 | |||
60 | Definition to_flocq_round_mode (m : round_mode) : BinarySingleNaN.mode := | ||
61 | match m with Floor => mode_DN | Ceil => mode_UP | NearestEven => mode_NE end. | ||
62 | |||
63 | Definition round (m : round_mode) : float → float := | ||
64 | Bnearbyint prec emax (refl_equal _) (λ _, indef_nan) (to_flocq_round_mode m). | ||
65 | |||
66 | (* For add: not [mode_DN]; otherwise [+0.0 + -0.0] would yield [-0.0], but | ||
67 | [inf / (+0.0 + -0.0)] yields [inf] in C++, so this cannot be the case. *) | ||
68 | (* C++ compiles floating point addition to the x86 ADDSD instruction. Looking | ||
69 | at the Intel x86 Software Developer's Manual, it seems that the default | ||
70 | rounding mode on x86 is Round to Nearest (even); see table 4-8. (In §4.8.4.) *) | ||
71 | Definition add : float → float → float := | ||
72 | Bplus _ _ Hprec Hprec_emax (λ _ _, indef_nan) mode_NE. | ||
73 | Definition sub : float → float → float := | ||
74 | Bminus _ _ Hprec Hprec_emax (λ _ _, indef_nan) mode_NE. | ||
75 | Definition mul : float → float → float := | ||
76 | Bmult _ _ Hprec Hprec_emax (λ _ _, indef_nan) mode_NE. | ||
77 | Definition div : float → float → float := | ||
78 | Bdiv _ _ Hprec Hprec_emax (λ _ _, indef_nan) mode_NE. | ||
79 | |||
80 | Definition eqb (f1 f2 : float) : bool := | ||
81 | bool_decide (b64_compare f1 f2 = Some Eq). | ||
82 | |||
83 | Definition ltb (f1 f2 : float) : bool := | ||
84 | bool_decide (b64_compare f1 f2 = Some Lt). | ||
85 | End Float. | ||