aboutsummaryrefslogtreecommitdiffstats
path: root/mppa_k1c/Asm.v
diff options
context:
space:
mode:
authorCyril SIX <cyril.six@kalray.eu>2019-02-12 15:10:20 +0100
committerCyril SIX <cyril.six@kalray.eu>2019-02-12 15:10:20 +0100
commit40bc8bf185c5cd3c5620cf4fe24ebcc9511c79fb (patch)
tree24d61f0ca190e6416bdd9081f51ae36ee4403de0 /mppa_k1c/Asm.v
parent1f38df638dfd4ae2aa85467cc28f6b4fa1da03b1 (diff)
downloadcompcert-kvx-40bc8bf185c5cd3c5620cf4fe24ebcc9511c79fb.tar.gz
compcert-kvx-40bc8bf185c5cd3c5620cf4fe24ebcc9511c79fb.zip
Added Ofloatconst and Osingleconst (not integrated in scheduler yet)
Diffstat (limited to 'mppa_k1c/Asm.v')
-rw-r--r--mppa_k1c/Asm.v12
1 files changed, 12 insertions, 0 deletions
diff --git a/mppa_k1c/Asm.v b/mppa_k1c/Asm.v
index 0ce3b455..b22ea100 100644
--- a/mppa_k1c/Asm.v
+++ b/mppa_k1c/Asm.v
@@ -105,6 +105,12 @@ Inductive instruction : Type :=
(** Arith RI64 *)
| Pmakel (rd: ireg) (imm: int64) (**r load immediate long *)
+ (** Arith RF32 *)
+ | Pmakefs (rd: ireg) (imm: float32)
+
+ (** Arith RF64 *)
+ | Pmakef (rd: ireg) (imm: float)
+
(** Arith RRR *)
| Pcompw (it: itest) (rd rs1 rs2: ireg) (**r comparison word *)
| Pcompl (it: itest) (rd rs1 rs2: ireg) (**r comparison long *)
@@ -195,6 +201,12 @@ Definition basic_to_instruction (b: basic) :=
(* RI64 *)
| PArithRI64 Asmblock.Pmakel rd imm => Pmakel rd imm
+ (* RF32 *)
+ | PArithRF32 Asmblock.Pmakefs rd imm => Pmakefs rd imm
+
+ (* RF64 *)
+ | PArithRF64 Asmblock.Pmakef rd imm => Pmakef rd imm
+
(* RRR *)
| PArithRRR (Asmblock.Pcompw it) rd rs1 rs2 => Pcompw it rd rs1 rs2
| PArithRRR (Asmblock.Pcompl it) rd rs1 rs2 => Pcompl it rd rs1 rs2