aboutsummaryrefslogtreecommitdiffstats
path: root/powerpc/Machregs.v
diff options
context:
space:
mode:
authorBernhard Schommer <bernhardschommer@gmail.com>2015-09-09 09:47:57 +0200
committerBernhard Schommer <bernhardschommer@gmail.com>2015-09-09 09:47:57 +0200
commit62277541b0ea1c687c64df79a543f776b9f58f29 (patch)
tree018b97b58a96f6c778db85ed16b382e295efee31 /powerpc/Machregs.v
parentfc4af0aa2601b290ec62dceef11599171f056770 (diff)
downloadcompcert-62277541b0ea1c687c64df79a543f776b9f58f29.tar.gz
compcert-62277541b0ea1c687c64df79a543f776b9f58f29.zip
Added an builtin for the atomic exchange operation.
The new builtin __builtin_atomic_exchange(int *a, int *b, int *c) stores *b in *a and sets *c to the old value of *a.
Diffstat (limited to 'powerpc/Machregs.v')
-rw-r--r--powerpc/Machregs.v8
1 files changed, 7 insertions, 1 deletions
diff --git a/powerpc/Machregs.v b/powerpc/Machregs.v
index a2017dca..f07d488b 100644
--- a/powerpc/Machregs.v
+++ b/powerpc/Machregs.v
@@ -160,9 +160,14 @@ Fixpoint destroyed_by_clobber (cl: list string): list mreg :=
end
end.
+Definition builtin_atomic_exchange := ident_of_string "__builtin_atomic_exchange".
+
Definition destroyed_by_builtin (ef: external_function): list mreg :=
match ef with
- | EF_builtin _ _ => F13 :: nil
+ | EF_builtin id sg =>
+ if ident_eq id builtin_atomic_exchange then R10::R11::F13:: nil
+ else
+ F13 :: nil
| EF_vload _ => R11 :: nil
| EF_vstore Mint64 => R10 :: R11 :: R12 :: nil
| EF_vstore _ => R11 :: R12 :: nil
@@ -183,6 +188,7 @@ Definition temp_for_parent_frame: mreg :=
Definition mregs_for_operation (op: operation): list (option mreg) * option mreg :=
(nil, None).
+
Definition mregs_for_builtin (ef: external_function): list (option mreg) * list (option mreg) :=
(nil, nil).