aboutsummaryrefslogtreecommitdiffstats
path: root/aarch64/Asm.v
diff options
context:
space:
mode:
Diffstat (limited to 'aarch64/Asm.v')
-rw-r--r--aarch64/Asm.v28
1 files changed, 21 insertions, 7 deletions
diff --git a/aarch64/Asm.v b/aarch64/Asm.v
index dc1f025f..5f109224 100644
--- a/aarch64/Asm.v
+++ b/aarch64/Asm.v
@@ -201,8 +201,8 @@ Inductive instruction: Type :=
| Pstrx_a (rs: ireg) (a: addressing) (**r store int64 as any64 *)
| Pstrb (rs: ireg) (a: addressing) (**r store int8 *)
| Pstrh (rs: ireg) (a: addressing) (**r store int16 *)
- | Pstpw (rs1 rs2: ireg) (chk1 chk2: memory_chunk) (a: addressing) (**r store two int64 *)
- | Pstpx (rs1 rs2: ireg) (chk1 chk2: memory_chunk) (a: addressing) (**r store two int64 *)
+ | Pstpw (rs1 rs2: ireg) (chk1 chk2: memory_chunk) (a: addressing) (**r store two int64 *)
+ | Pstpx (rs1 rs2: ireg) (chk1 chk2: memory_chunk) (a: addressing) (**r store two int64 *)
(** Integer arithmetic, immediate *)
| Paddimm (sz: isize) (rd: iregsp) (r1: iregsp) (n: Z) (**r addition *)
| Psubimm (sz: isize) (rd: iregsp) (r1: iregsp) (n: Z) (**r subtraction *)
@@ -273,9 +273,13 @@ Inductive instruction: Type :=
| Pldrs (rd: freg) (a: addressing) (**r load float32 (single precision) *)
| Pldrd (rd: freg) (a: addressing) (**r load float64 (double precision) *)
| Pldrd_a (rd: freg) (a: addressing) (**r load float64 as any64 *)
+ | Pldps (rd1 rd2: freg) (chk1 chk2: memory_chunk) (a: addressing) (**r load two float32 *)
+ | Pldpd (rd1 rd2: freg) (chk1 chk2: memory_chunk) (a: addressing) (**r load two float64 *)
| Pstrs (rs: freg) (a: addressing) (**r store float32 *)
| Pstrd (rs: freg) (a: addressing) (**r store float64 *)
| Pstrd_a (rs: freg) (a: addressing) (**r store float64 as any64 *)
+ | Pstps (rd1 rd2: freg) (chk1 chk2: memory_chunk) (a: addressing) (**r store two float32 *)
+ | Pstpd (rd1 rd2: freg) (chk1 chk2: memory_chunk) (a: addressing) (**r store two float64 *)
(** Floating-point move *)
| Pfmov (rd r1: freg)
| Pfmovimms (rd: freg) (f: float32) (**r load float32 constant *)
@@ -798,7 +802,7 @@ Definition exec_load_double (chk1 chk2: memory_chunk) (transf: val -> val)
(a: addressing) (rd1 rd2: preg) (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
@@ -824,7 +828,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
@@ -1258,17 +1262,27 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out
| _ => Stuck
end
| Pbuiltin ef args res => Stuck (**r treated specially below *)
- (** The following instructions and directives are not generated directly
- by Asmgen, so we do not model them. *)
+ (** loads and stores pairs int/int64 *)
| Pldpw rd1 rd2 chk1 chk2 a =>
exec_load_double chk1 chk2 (fun v => v) a rd1 rd2 rs m
| Pldpx rd1 rd2 chk1 chk2 a =>
exec_load_double chk1 chk2 (fun v => v) a rd1 rd2 rs m
- | Pnop => Next (nextinstr rs) m
| Pstpw rs1 rs2 chk1 chk2 a =>
exec_store_double chk1 chk2 a rs#rs1 rs#rs2 rs m
| Pstpx rs1 rs2 chk1 chk2 a =>
exec_store_double chk1 chk2 a rs#rs1 rs#rs2 rs m
+ (** loads and stores pairs floating-point *)
+ | Pldps rd1 rd2 chk1 chk2 a =>
+ exec_load_double chk1 chk2 (fun v => v) a rd1 rd2 rs m
+ | Pldpd rd1 rd2 chk1 chk2 a =>
+ exec_load_double chk1 chk2 (fun v => v) a rd1 rd2 rs m
+ | Pstps rs1 rs2 chk1 chk2 a =>
+ exec_store_double chk1 chk2 a rs#rs1 rs#rs2 rs m
+ | Pstpd rs1 rs2 chk1 chk2 a =>
+ exec_store_double chk1 chk2 a rs#rs1 rs#rs2 rs m
+ | Pnop => Next (nextinstr rs) m
+ (** The following instructions and directives are not generated directly
+ by Asmgen, so we do not model them. *)
| Pcls _ _ _
| Pclz _ _ _
| Prev _ _ _