aboutsummaryrefslogtreecommitdiffstats
path: root/kvx/Asmvliw.v
diff options
context:
space:
mode:
authorDavid Monniaux <David.Monniaux@univ-grenoble-alpes.fr>2021-12-12 17:31:41 +0100
committerDavid Monniaux <David.Monniaux@univ-grenoble-alpes.fr>2021-12-12 17:31:41 +0100
commit4f7d6d6a081de52fe1151a29d44221f4fc35f7be (patch)
treedc5cd69f25cb0b42ef906aed1848b1b0bc9fbd5c /kvx/Asmvliw.v
parent1ecddb62bc5aa8e6a1f6c1a1a2da2e2a8e2b100f (diff)
downloadcompcert-kvx-4f7d6d6a081de52fe1151a29d44221f4fc35f7be.tar.gz
compcert-kvx-4f7d6d6a081de52fe1151a29d44221f4fc35f7be.zip
Asm level
Diffstat (limited to 'kvx/Asmvliw.v')
-rw-r--r--kvx/Asmvliw.v9
1 files changed, 8 insertions, 1 deletions
diff --git a/kvx/Asmvliw.v b/kvx/Asmvliw.v
index 45b230e6..b3c0c8fa 100644
--- a/kvx/Asmvliw.v
+++ b/kvx/Asmvliw.v
@@ -402,7 +402,10 @@ Inductive arith_name_rr : Type :=
| Pfixedudrzz (**r Integer conversion from floating point (float -> unsigned long) *)
| Pfixeddrzz_i32 (**r Integer conversion from floating point (float -> int) *)
| Pfixedudrzz_i32 (**r Integer conversion from floating point (float -> unsigned int) *)
-.
+ | Pfixedwrne (**r Integer conversion from floating point *)
+ | Pfixeduwrne (**r Integer conversion from floating point (f32 -> 32 bits unsigned *)
+ | Pfixeddrne (**r Integer conversion from floating point (i64 -> 64 bits) *)
+ | Pfixedudrne. (**r unsigned Integer conversion from floating point (u64 -> 64 bits) *)
Inductive arith_name_ri32 : Type :=
| Pmake (**r load immediate *)
@@ -955,6 +958,10 @@ Definition arith_eval_rr n v :=
| Pfixedudrzz => Val.maketotal (Val.longuoffloat v)
| Pfixeddrzz_i32 => Val.maketotal (Val.intoffloat v)
| Pfixedudrzz_i32 => Val.maketotal (Val.intuoffloat v)
+ | Pfixedudrne => Val.maketotal (Val.longuoffloat_ne v)
+ | Pfixeddrne => Val.maketotal (Val.longoffloat_ne v)
+ | Pfixeduwrne => Val.maketotal (Val.intuoffloat_ne v)
+ | Pfixedwrne => Val.maketotal (Val.intoffloat_ne v)
end.
Definition arith_eval_ri32 n i :=