aboutsummaryrefslogtreecommitdiffstats
path: root/powerpc/Asmgen.v
diff options
context:
space:
mode:
authorxleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2013-03-16 16:51:42 +0000
committerxleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2013-03-16 16:51:42 +0000
commit707b6a1ae9660b13cf6f68c7c0ce74017f5981c5 (patch)
tree166a21d507612d60db40737333ddec5003fc81aa /powerpc/Asmgen.v
parente44df4563f1cb893ab25b2a8b25d5b83095205be (diff)
downloadcompcert-kvx-707b6a1ae9660b13cf6f68c7c0ce74017f5981c5.tar.gz
compcert-kvx-707b6a1ae9660b13cf6f68c7c0ce74017f5981c5.zip
Assorted changes to reduce stack and heap requirements when compiling very big functions.
git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@2151 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'powerpc/Asmgen.v')
-rw-r--r--powerpc/Asmgen.v28
1 files changed, 23 insertions, 5 deletions
diff --git a/powerpc/Asmgen.v b/powerpc/Asmgen.v
index 0035dff8..39b84e0b 100644
--- a/powerpc/Asmgen.v
+++ b/powerpc/Asmgen.v
@@ -621,7 +621,7 @@ Definition transl_instr (f: Mach.function) (i: Mach.instruction)
(** Translation of a code sequence *)
-Definition r11_is_parent (before: bool) (i: Mach.instruction) : bool :=
+Definition it1_is_parent (before: bool) (i: Mach.instruction) : bool :=
match i with
| Msetstack src ofs ty => before
| Mgetparam ofs ty dst => negb (mreg_eq dst IT1)
@@ -629,21 +629,39 @@ Definition r11_is_parent (before: bool) (i: Mach.instruction) : bool :=
| _ => false
end.
-Fixpoint transl_code (f: Mach.function) (il: list Mach.instruction) (r11p: bool) :=
+(** This is the naive definition that we no longer use because it
+ is not tail-recursive. It is kept as specification. *)
+
+Fixpoint transl_code (f: Mach.function) (il: list Mach.instruction) (it1p: bool) :=
match il with
| nil => OK nil
| i1 :: il' =>
- do k <- transl_code f il' (r11_is_parent r11p i1);
- transl_instr f i1 r11p k
+ do k <- transl_code f il' (it1_is_parent it1p i1);
+ transl_instr f i1 it1p k
end.
+(** This is an equivalent definition in continuation-passing style
+ that runs in constant stack space. *)
+
+Fixpoint transl_code_rec (f: Mach.function) (il: list Mach.instruction)
+ (it1p: bool) (k: code -> res code) :=
+ match il with
+ | nil => k nil
+ | i1 :: il' =>
+ transl_code_rec f il' (it1_is_parent it1p i1)
+ (fun c1 => do c2 <- transl_instr f i1 it1p c1; k c2)
+ end.
+
+Definition transl_code' (f: Mach.function) (il: list Mach.instruction) (it1p: bool) :=
+ transl_code_rec f il it1p (fun c => OK c).
+
(** Translation of a whole function. Note that we must check
that the generated code contains less than [2^32] instructions,
otherwise the offset part of the [PC] code pointer could wrap
around, leading to incorrect executions. *)
Definition transl_function (f: Mach.function) :=
- do c <- transl_code f f.(Mach.fn_code) false;
+ do c <- transl_code' f f.(Mach.fn_code) false;
OK (Pallocframe f.(fn_stacksize) f.(fn_link_ofs) ::
Pmflr GPR0 ::
Pstw GPR0 (Cint f.(fn_retaddr_ofs)) GPR1 :: c).