aboutsummaryrefslogtreecommitdiffstats
path: root/arm/Asm.v
diff options
context:
space:
mode:
authorxleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2010-09-02 12:42:19 +0000
committerxleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2010-09-02 12:42:19 +0000
commit265fa07b34a813ba9d8249ddad82d71e6002c10d (patch)
tree45831b1793c7920b10969fc7cf6316c202d78e91 /arm/Asm.v
parent94470fb6a652cb993982269fcb7a0e8319b54488 (diff)
downloadcompcert-kvx-265fa07b34a813ba9d8249ddad82d71e6002c10d.tar.gz
compcert-kvx-265fa07b34a813ba9d8249ddad82d71e6002c10d.zip
Merge of the reuse-temps branch:
- Reload temporaries are marked as destroyed (set to Vundef) across operations in the semantics of LTL, LTLin, Linear and Mach, allowing Asmgen to reuse them. - Added IA32 port. - Cleaned up float conversions and axiomatization of floats. git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@1499 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'arm/Asm.v')
-rw-r--r--arm/Asm.v13
1 files changed, 5 insertions, 8 deletions
diff --git a/arm/Asm.v b/arm/Asm.v
index d160be78..7ea1a8a3 100644
--- a/arm/Asm.v
+++ b/arm/Asm.v
@@ -152,9 +152,7 @@ Inductive instruction : Type :=
| Pcmf: freg -> freg -> instruction (**r float comparison *)
| Pdvfd: freg -> freg -> freg -> instruction (**r float division *)
| Pfixz: ireg -> freg -> instruction (**r float to signed int *)
- | Pfixzu: ireg -> freg -> instruction (**r float to unsigned int *)
| Pfltd: freg -> ireg -> instruction (**r signed int to float *)
- | Pfltud: freg -> ireg -> instruction (**r unsigned int to float *)
| Pldfd: freg -> ireg -> int -> instruction (**r float64 load *)
| Pldfs: freg -> ireg -> int -> instruction (**r float32 load *)
| Plifd: freg -> float -> instruction (**r load float constant *)
@@ -470,12 +468,8 @@ Definition exec_instr (c: code) (i: instruction) (rs: regset) (m: mem) : outcome
OK (nextinstr (rs#r1 <- (Val.divf rs#r2 rs#r3))) m
| Pfixz r1 r2 =>
OK (nextinstr (rs#r1 <- (Val.intoffloat rs#r2))) m
- | Pfixzu r1 r2 =>
- OK (nextinstr (rs#r1 <- (Val.intuoffloat rs#r2))) m
| Pfltd r1 r2 =>
OK (nextinstr (rs#r1 <- (Val.floatofint rs#r2))) m
- | Pfltud r1 r2 =>
- OK (nextinstr (rs#r1 <- (Val.floatofintu rs#r2))) m
| Pldfd r1 r2 n =>
exec_load Mfloat64 (Val.add rs#r2 (Vint n)) r1 rs m
| Pldfs r1 r2 n =>
@@ -492,8 +486,11 @@ Definition exec_instr (c: code) (i: instruction) (rs: regset) (m: mem) : outcome
OK (nextinstr (rs#r1 <- (Val.mulf rs#r2 rs#r3))) m
| Pstfd r1 r2 n =>
exec_store Mfloat64 (Val.add rs#r2 (Vint n)) r1 rs m
- | Pstfs r1 r2 n =>
- exec_store Mfloat32 (Val.add rs#r2 (Vint n)) r1 rs m
+ | Pstfs r1 r2 n =>
+ match exec_store Mfloat32 (Val.add rs#r2 (Vint n)) r1 rs m with
+ | OK rs' m' => OK (rs'#FR3 <- Vundef) m'
+ | Error => Error
+ end
| Psufd r1 r2 r3 =>
OK (nextinstr (rs#r1 <- (Val.subf rs#r2 rs#r3))) m
(* Pseudo-instructions *)