aboutsummaryrefslogtreecommitdiffstats
path: root/aarch64/Asmblock.v
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2021-01-22 14:17:39 +0100
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2021-01-22 14:17:39 +0100
commitcac0556d551438b5dc3bbb6b701ae0101b214e0f (patch)
treec97b791ac6c8029f512e9b110dd5e430817d806d /aarch64/Asmblock.v
parente265be77756b14b1d830a0d0faf1b105494bbb43 (diff)
parent0bc5b0f9fe8a2463ddb147671359a5b374cfd50c (diff)
downloadcompcert-kvx-cac0556d551438b5dc3bbb6b701ae0101b214e0f.tar.gz
compcert-kvx-cac0556d551438b5dc3bbb6b701ae0101b214e0f.zip
Merge remote-tracking branch 'origin/aarch64-peephole' into kvx-work
Diffstat (limited to 'aarch64/Asmblock.v')
-rw-r--r--aarch64/Asmblock.v12
1 files changed, 8 insertions, 4 deletions
diff --git a/aarch64/Asmblock.v b/aarch64/Asmblock.v
index ed84b7d8..c606002a 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,13 @@ Inductive store_rs_a : Type :=
Inductive store_rs1_rs2_a : Type :=
| Pstpw
| Pstpx
+ | Pstps
+ | Pstpd
.
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 +485,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 +509,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