aboutsummaryrefslogtreecommitdiffstats
path: root/aarch64/Asmblock.v
diff options
context:
space:
mode:
authorLéo Gourdin <leo.gourdin@univ-grenoble-alpes.fr>2021-01-20 16:46:52 +0100
committerLéo Gourdin <leo.gourdin@univ-grenoble-alpes.fr>2021-01-20 16:46:52 +0100
commit8781c27a16463fd8f83a9c6eaba7b76846bd296c (patch)
tree17beca9da05a776da1193eb9c2002ea79a3261b5 /aarch64/Asmblock.v
parent87e268bc9e6079b5aea31cd6c20e30d00adf2bb8 (diff)
downloadcompcert-kvx-8781c27a16463fd8f83a9c6eaba7b76846bd296c.tar.gz
compcert-kvx-8781c27a16463fd8f83a9c6eaba7b76846bd296c.zip
Adding fp loads pair
Diffstat (limited to 'aarch64/Asmblock.v')
-rw-r--r--aarch64/Asmblock.v11
1 files changed, 7 insertions, 4 deletions
diff --git a/aarch64/Asmblock.v b/aarch64/Asmblock.v
index ed84b7d8..f5d2abbe 100644
--- a/aarch64/Asmblock.v
+++ b/aarch64/Asmblock.v
@@ -110,11 +110,13 @@ Inductive load_rd_a: Type :=
Inductive load_rd1_rd2_a: Type :=
| Pldpw
| Pldpx
+ | Pldps
+ | Pldpd
.
Inductive ld_instruction: Type :=
| PLd_rd_a (ld: load_rd_a) (rd: dreg) (a: addressing)
- | Pldp (ld: load_rd1_rd2_a) (rd1 rd2: ireg) (chk1 chk2: memory_chunk) (a: addressing) (**r load two int64 *)
+ | Pldp (ld: load_rd1_rd2_a) (rd1 rd2: dreg) (chk1 chk2: memory_chunk) (a: addressing) (**r load two int64 *)
.
Inductive store_rs_a : Type :=
@@ -134,11 +136,12 @@ Inductive store_rs_a : Type :=
Inductive store_rs1_rs2_a : Type :=
| Pstpw
| Pstpx
+ (* TODO *)
.
Inductive st_instruction : Type :=
| PSt_rs_a (st: store_rs_a) (rs: dreg) (a: addressing)
- | Pstp (st: store_rs1_rs2_a) (rs1 rs2: ireg) (chk1 chk2: memory_chunk) (a: addressing) (**r store two int64 *)
+ | Pstp (st: store_rs1_rs2_a) (rs1 rs2: dreg) (chk1 chk2: memory_chunk) (a: addressing) (**r store two int64 *)
.
Inductive arith_p : Type :=
@@ -481,7 +484,7 @@ Definition exec_load_double (chk1 chk2: memory_chunk) (transf: val -> val)
(a: addressing) (rd1 rd2: dreg) (rs: regset) (m: mem) :=
if is_pair_addressing_mode_correct a then
let addr := (eval_addressing a rs) in
- let ofs := match chk1 with | Mint32 | Many32 => 4 | _ => 8 end in
+ let ofs := match chk1 with | Mint32 | Mfloat32 | Many32 => 4 | _ => 8 end in
let addr' := (eval_addressing (get_offset_addr a ofs) rs) in
match Mem.loadv chk1 m addr with
| None => Stuck
@@ -505,7 +508,7 @@ Definition exec_store_double (chk1 chk2: memory_chunk)
(rs: regset) (m: mem) :=
if is_pair_addressing_mode_correct a then
let addr := (eval_addressing a rs) in
- let ofs := match chk1 with | Mint32 | Many32 => 4 | _ => 8 end in
+ let ofs := match chk1 with | Mint32 | Mfloat32 | Many32 => 4 | _ => 8 end in
let addr' := (eval_addressing (get_offset_addr a ofs) rs) in
match Mem.storev chk1 m addr v1 with
| None => Stuck