summaryrefslogtreecommitdiff
path: root/backend/Stacking.v
diff options
context:
space:
mode:
authorGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2008-04-12 15:06:13 +0000
committerGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2008-04-12 15:06:13 +0000
commit41ef4d52e3c328d930979115cb4fd388cda09440 (patch)
tree82cdf63d42d4231d19610a57af2a849cf6f1b0ad /backend/Stacking.v
parentaaa49526068f528f2233de0dace43549432fba52 (diff)
Revu le traitement de la 'red zone' en bas de la pile
git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@605 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'backend/Stacking.v')
-rw-r--r--backend/Stacking.v6
1 files changed, 4 insertions, 2 deletions
diff --git a/backend/Stacking.v b/backend/Stacking.v
index a1310aa..3f08daa 100644
--- a/backend/Stacking.v
+++ b/backend/Stacking.v
@@ -52,6 +52,8 @@ The [frame_env] compilation environment records the positions of
the boundaries between areas in the frame part.
*)
+Definition fe_ofs_arg := 24.
+
Record frame_env : Set := mk_frame_env {
fe_size: Z;
fe_ofs_link: Z;
@@ -68,7 +70,7 @@ Record frame_env : Set := mk_frame_env {
function. *)
Definition make_env (b: bounds) :=
- let oil := 4 * b.(bound_outgoing) in (* integer locals *)
+ let oil := 24 + 4 * b.(bound_outgoing) in (* integer locals *)
let oics := oil + 4 * b.(bound_int_local) in (* integer callee-saves *)
let oendi := oics + 4 * b.(bound_int_callee_save) in
let ofl := align oendi 8 in (* float locals *)
@@ -95,7 +97,7 @@ Definition offset_of_index (fe: frame_env) (idx: frame_index) :=
| FI_retaddr => fe.(fe_ofs_retaddr)
| FI_local x Tint => fe.(fe_ofs_int_local) + 4 * x
| FI_local x Tfloat => fe.(fe_ofs_float_local) + 8 * x
- | FI_arg x ty => 4 * x
+ | FI_arg x ty => fe_ofs_arg + 4 * x
| FI_saved_int x => fe.(fe_ofs_int_callee_save) + 4 * x
| FI_saved_float x => fe.(fe_ofs_float_callee_save) + 8 * x
end.