aboutsummaryrefslogtreecommitdiffstats
path: root/mppa_k1c/Archi.v
diff options
context:
space:
mode:
Diffstat (limited to 'mppa_k1c/Archi.v')
-rw-r--r--mppa_k1c/Archi.v42
1 files changed, 28 insertions, 14 deletions
diff --git a/mppa_k1c/Archi.v b/mppa_k1c/Archi.v
index bbe66c5b..69b32c7c 100644
--- a/mppa_k1c/Archi.v
+++ b/mppa_k1c/Archi.v
@@ -14,11 +14,10 @@
(* *)
(* *********************************************************************)
-(** Architecture-dependent parameters for RISC-V *)
+(** Architecture-dependent parameters for MPPA K1c. Mostly copied from the Risc-V backend *)
-Require Import ZArith.
-Require Import Fappli_IEEE.
-Require Import Fappli_IEEE_bits.
+Require Import ZArith List.
+Require Import Binary Bits.
Definition ptr64 := true.
@@ -34,6 +33,8 @@ Proof.
unfold splitlong. destruct ptr64; simpl; congruence.
Qed.
+(** THIS IS NOT CHECKED ! NONE OF THIS ! *)
+
(** Section 7.3: "Except when otherwise stated, if the result of a
floating-point operation is NaN, it is the canonical NaN. The
canonical NaN has a positive sign and all significand bits clear
@@ -41,23 +42,36 @@ Qed.
We need to extend the [choose_binop_pl] functions to account for
this case. *)
-Program Definition default_pl_64 : bool * nan_pl 53 :=
- (false, iter_nat 51 _ xO xH).
+Definition default_nan_64 := (false, iter_nat 51 _ xO xH).
+Definition default_nan_32 := (false, iter_nat 22 _ xO xH).
+
+(* Always choose the first NaN argument, if any *)
+
+Definition choose_nan_64 (l: list (bool * positive)) : bool * positive :=
+ match l with nil => default_nan_64 | n :: _ => n end.
+
+Definition choose_nan_32 (l: list (bool * positive)) : bool * positive :=
+ match l with nil => default_nan_32 | n :: _ => n end.
+
+Definition fpu_returns_default_qNaN := false.
-Definition choose_binop_pl_64 (s1: bool) (pl1: nan_pl 53) (s2: bool) (pl2: nan_pl 53) :=
- false. (**r always choose first NaN *)
+Lemma choose_nan_64_idem: forall n,
+ choose_nan_64 (n :: n :: nil) = choose_nan_64 (n :: nil).
+Proof. auto. Qed.
-Program Definition default_pl_32 : bool * nan_pl 24 :=
- (false, iter_nat 22 _ xO xH).
+Lemma choose_nan_32_idem: forall n,
+ choose_nan_32 (n :: n :: nil) = choose_nan_32 (n :: nil).
+Proof. auto. Qed.
-Definition choose_binop_pl_32 (s1: bool) (pl1: nan_pl 24) (s2: bool) (pl2: nan_pl 24) :=
- false. (**r always choose first NaN *)
+Definition fma_order {A: Type} (x y z: A) := (x, z, y).
+Definition fma_invalid_mul_is_nan := false.
Definition float_of_single_preserves_sNaN := false.
Global Opaque ptr64 big_endian splitlong
- default_pl_64 choose_binop_pl_64
- default_pl_32 choose_binop_pl_32
+ default_nan_64 choose_nan_64
+ default_nan_32 choose_nan_32
+ fma_order fma_invalid_mul_is_nan
float_of_single_preserves_sNaN.
(** Whether to generate position-independent code or not *)