aboutsummaryrefslogtreecommitdiffstats
path: root/mppa_k1c/Machregs.v
diff options
context:
space:
mode:
authorCyril SIX <cyril.six@kalray.eu>2018-04-20 10:31:56 +0200
committerCyril SIX <cyril.six@kalray.eu>2018-04-20 10:31:56 +0200
commitaa25ec270b651186154523ec71a3888b50994d70 (patch)
treea4572adf5cde4f97642885cb1f90f73a20d169cd /mppa_k1c/Machregs.v
parent41a048fa4bb9ddefd4e4acff2207251bb3ddbf06 (diff)
downloadcompcert-kvx-aa25ec270b651186154523ec71a3888b50994d70.tar.gz
compcert-kvx-aa25ec270b651186154523ec71a3888b50994d70.zip
MPPA - Oshrximm + Mgetparam + FP is GPR10 + bug
Added Oshrximm and Mgetparam -> mmult.c divide & conqueer generates FP is now GPR10 instead of being a mix of GPR30 and GPR32 Corrected a bug where Pgoto and Pj_l were given the same interpretation, where in fact there's a fundamental difference : Pgoto is supposed to have a function name (symbol), while Pj_l is supposed to have a label name (print_label). This led to having undefinite labels in the code.
Diffstat (limited to 'mppa_k1c/Machregs.v')
-rw-r--r--mppa_k1c/Machregs.v10
1 files changed, 6 insertions, 4 deletions
diff --git a/mppa_k1c/Machregs.v b/mppa_k1c/Machregs.v
index ce86a06f..ed582c98 100644
--- a/mppa_k1c/Machregs.v
+++ b/mppa_k1c/Machregs.v
@@ -42,7 +42,7 @@ Require Import Op.
Inductive mreg: Type :=
(* Allocatable General Purpose regs. *)
| R0 | R1 | R2 | R3 | R4 | R5 | R6 | R7 | R9
- (* R10 to R14 are reserved *) | R15 | R16 | R17 | R18 | R19
+ | R10 (* R11 to R14 res *) | R15 | R16 | R17 | R18 | R19
| R20 | R21 | R22 | R23 | R24 | R25 | R26 | R27 | R28 | R29
| R30 | R32 | R33 | R34 | R35 | R36 | R37 | R38 | R39
| R40 | R41 | R42 | R43 | R44 | R45 | R46 | R47 | R48 | R49
@@ -55,7 +55,7 @@ Global Opaque mreg_eq.
Definition all_mregs :=
R0 :: R1 :: R2 :: R3 :: R4 :: R5 :: R6 :: R7 :: R9
- :: R15 :: R16 :: R17 :: R18 :: R19
+ :: R10 :: R15 :: R16 :: R17 :: R18 :: R19
:: R20 :: R21 :: R22 :: R23 :: R24 :: R25 :: R26 :: R27 :: R28 :: R29
:: R30 :: R32 :: R33 :: R34 :: R35 :: R36 :: R37 :: R38 :: R39
:: R40 :: R41 :: R42 :: R43 :: R44 :: R45 :: R46 :: R47 :: R48 :: R49
@@ -87,6 +87,7 @@ Module IndexedMreg <: INDEXED_TYPE.
match r with
| R0 => 1 | R1 => 2 | R2 => 3 | R3 => 4 | R4 => 5
| R5 => 6 | R6 => 7 | R7 => 8 | R9 => 10
+ | R10 => 11
| R15 => 16 | R16 => 17 | R17 => 18 | R18 => 19 | R19 => 20
| R20 => 21 | R21 => 22 | R22 => 23 | R23 => 24 | R24 => 25
| R25 => 26 | R26 => 27 | R27 => 28 | R28 => 29 | R29 => 30
@@ -115,6 +116,7 @@ Local Open Scope string_scope.
Definition register_names :=
("R0" , R0) :: ("R1" , R1) :: ("R2" , R2) :: ("R3" , R3) :: ("R4" , R4)
:: ("R5" , R5) :: ("R6" , R6) :: ("R7" , R7) :: ("R9" , R9)
+ :: ("R10", R10)
:: ("R15", R15) :: ("R16", R16) :: ("R17", R17) :: ("R18", R18) :: ("R19", R19)
:: ("R20", R20) :: ("R21", R21) :: ("R22", R22) :: ("R23", R23) :: ("R24", R24)
:: ("R25", R25) :: ("R26", R26) :: ("R27", R27) :: ("R28", R28) :: ("R29", R29)
@@ -173,9 +175,9 @@ Definition destroyed_by_builtin (ef: external_function): list mreg :=
Definition destroyed_by_setstack (ty: typ): list mreg := nil.
-Definition destroyed_at_function_entry: list mreg := R32 :: nil.
+Definition destroyed_at_function_entry: list mreg := R10 :: nil.
-Definition temp_for_parent_frame: mreg := R9. (* FIXME - and R8 ?? *)
+Definition temp_for_parent_frame: mreg := R10. (* FIXME - and R8 ?? *)
Definition destroyed_at_indirect_call: list mreg := nil.
(* R10 :: R11 :: R12 :: R13 :: R14 :: R15 :: R16 :: R17 :: nil. *)