aboutsummaryrefslogtreecommitdiffstats
path: root/x86/Conventions1.v
diff options
context:
space:
mode:
Diffstat (limited to 'x86/Conventions1.v')
-rw-r--r--x86/Conventions1.v17
1 files changed, 17 insertions, 0 deletions
diff --git a/x86/Conventions1.v b/x86/Conventions1.v
index a4e3b970..abd22001 100644
--- a/x86/Conventions1.v
+++ b/x86/Conventions1.v
@@ -82,6 +82,23 @@ Definition dummy_float_reg := X0. (**r Used in [Regalloc]. *)
Definition callee_save_type := mreg_type.
+(** How to use registers for register allocation.
+ We favor the use of caller-save registers, using callee-save registers
+ only when no caller-save is available. *)
+
+Record alloc_regs := mk_alloc_regs {
+ preferred_int_regs: list mreg;
+ remaining_int_regs: list mreg;
+ preferred_float_regs: list mreg;
+ remaining_float_regs: list mreg
+}.
+
+Definition allocatable_registers (_: unit) :=
+ {| preferred_int_regs := int_caller_save_regs;
+ remaining_int_regs := int_callee_save_regs;
+ preferred_float_regs := float_caller_save_regs;
+ remaining_float_regs := float_callee_save_regs |}.
+
(** * Function calling conventions *)
(** The functions in this section determine the locations (machine registers