aboutsummaryrefslogtreecommitdiffstats
path: root/arm/Machregs.v
diff options
context:
space:
mode:
authorBernhard Schommer <bernhardschommer@gmail.com>2016-09-15 12:32:26 +0200
committerBernhard Schommer <bernhardschommer@gmail.com>2016-09-15 12:37:02 +0200
commit6f2b5713f8e378e6e074f35a537e86a497c64e35 (patch)
tree87d80f8a4dd0de1f0594fc67c48ff26a58c68056 /arm/Machregs.v
parent9124c9231c11effae6e32d73c6c8af7c4032f928 (diff)
downloadcompcert-kvx-6f2b5713f8e378e6e074f35a537e86a497c64e35.tar.gz
compcert-kvx-6f2b5713f8e378e6e074f35a537e86a497c64e35.zip
Add interference for indirect calls.
Avoids problems with overwritting the registe containing the function address. Bug 19779
Diffstat (limited to 'arm/Machregs.v')
-rw-r--r--arm/Machregs.v6
1 files changed, 5 insertions, 1 deletions
diff --git a/arm/Machregs.v b/arm/Machregs.v
index b43f9be6..e97df790 100644
--- a/arm/Machregs.v
+++ b/arm/Machregs.v
@@ -58,7 +58,7 @@ Proof.
Qed.
Instance Decidable_eq_mreg : forall (x y: mreg), Decidable (eq x y) := Decidable_eq mreg_eq.
-
+
Instance Finite_mreg : Finite mreg := {
Finite_elements := all_mregs;
Finite_elements_spec := all_mregs_complete
@@ -161,6 +161,9 @@ Definition destroyed_by_setstack (ty: typ): list mreg := nil.
Definition destroyed_at_function_entry: list mreg :=
R12 :: nil.
+Definition destroyed_at_indirect_call: list mreg :=
+ R0 :: R1 :: R2 :: R3 :: nil.
+
Definition temp_for_parent_frame: mreg :=
R12.
@@ -177,6 +180,7 @@ Global Opaque
destroyed_by_op destroyed_by_load destroyed_by_store
destroyed_by_cond destroyed_by_jumptable destroyed_by_builtin
destroyed_by_setstack destroyed_at_function_entry temp_for_parent_frame
+ destroyed_at_indirect_call
mregs_for_operation mregs_for_builtin.
(** Two-address operations. Return [true] if the first argument and