aboutsummaryrefslogtreecommitdiffstats
path: root/arm/Asm.v
diff options
context:
space:
mode:
authorXavier Leroy <xavier.leroy@college-de-france.fr>2020-12-06 16:47:01 +0100
committerXavier Leroy <xavier.leroy@college-de-france.fr>2020-12-06 16:47:01 +0100
commita4cfb9c2ffdef07fa0d478e66f279687c9823d42 (patch)
tree722f5c83886909ad2714720715c2b11ca88ac5f8 /arm/Asm.v
parent0df99dc46209a9fe5026b83227ef73280f0dab70 (diff)
downloadcompcert-kvx-a4cfb9c2ffdef07fa0d478e66f279687c9823d42.tar.gz
compcert-kvx-a4cfb9c2ffdef07fa0d478e66f279687c9823d42.zip
ARM modeling of registers destroyed by pseudo-instructions
Pflid destroys IR14 Inlined built-in functions destroy IR14
Diffstat (limited to 'arm/Asm.v')
-rw-r--r--arm/Asm.v4
1 files changed, 2 insertions, 2 deletions
diff --git a/arm/Asm.v b/arm/Asm.v
index 194074ac..293df274 100644
--- a/arm/Asm.v
+++ b/arm/Asm.v
@@ -696,7 +696,7 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out
| Pfsubd r1 r2 r3 =>
Next (nextinstr (rs#r1 <- (Val.subf rs#r2 rs#r3))) m
| Pflid r1 f =>
- Next (nextinstr (rs#r1 <- (Vfloat f))) m
+ Next (nextinstr (rs#IR14 <- Vundef #r1 <- (Vfloat f))) m
| Pfcmpd r1 r2 =>
Next (nextinstr (compare_float rs rs#r1 rs#r2)) m
| Pfcmpzd r1 =>
@@ -923,7 +923,7 @@ Inductive step: state -> trace -> state -> Prop :=
external_call ef ge vargs m t vres m' ->
rs' = nextinstr
(set_res res vres
- (undef_regs (map preg_of (destroyed_by_builtin ef)) rs)) ->
+ (undef_regs (IR IR14 :: map preg_of (destroyed_by_builtin ef)) rs)) ->
step (State rs m) t (State rs' m')
| exec_step_external:
forall b ef args res rs m t rs' m',