summaryrefslogtreecommitdiff
path: root/ia32
diff options
context:
space:
mode:
authorGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2013-03-16 16:51:42 +0000
committerGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2013-03-16 16:51:42 +0000
commit707b6a1ae9660b13cf6f68c7c0ce74017f5981c5 (patch)
tree166a21d507612d60db40737333ddec5003fc81aa /ia32
parente44df4563f1cb893ab25b2a8b25d5b83095205be (diff)
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 'ia32')
-rw-r--r--ia32/Asmgen.v28
-rw-r--r--ia32/Asmgenproof.v9
-rw-r--r--ia32/PrintOp.ml1
3 files changed, 29 insertions, 9 deletions
diff --git a/ia32/Asmgen.v b/ia32/Asmgen.v
index a7a629b..6e3ccf8 100644
--- a/ia32/Asmgen.v
+++ b/ia32/Asmgen.v
@@ -528,28 +528,46 @@ Definition transl_instr (f: Mach.function) (i: Mach.instruction)
(** Translation of a code sequence *)
-Definition edx_preserved (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)
| _ => false
end.
-Fixpoint transl_code (f: Mach.function) (il: list Mach.instruction) (edx_is_parent: 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' (edx_preserved edx_is_parent i1);
- transl_instr f i1 edx_is_parent 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 transf_function (f: Mach.function) : res Asm.code :=
- do c <- transl_code f f.(Mach.fn_code) true;
+ do c <- transl_code' f f.(Mach.fn_code) true;
if zlt (list_length_z c) Int.max_unsigned
then OK (Pallocframe f.(fn_stacksize) f.(fn_retaddr_ofs) f.(fn_link_ofs) :: c)
else Error (msg "code size exceeded").
diff --git a/ia32/Asmgenproof.v b/ia32/Asmgenproof.v
index 060d018..e43392a 100644
--- a/ia32/Asmgenproof.v
+++ b/ia32/Asmgenproof.v
@@ -332,7 +332,7 @@ Lemma transl_find_label:
end.
Proof.
intros. monadInv H. destruct (zlt (list_length_z x) Int.max_unsigned); inv EQ0.
- simpl. eapply transl_code_label; eauto.
+ simpl. eapply transl_code_label; eauto. rewrite transl_code'_transl_code in EQ; eauto.
Qed.
End TRANSL_LABEL.
@@ -375,6 +375,7 @@ Proof.
destruct i; try (intros [A B]; apply A). intros. subst c0. repeat constructor.
- intros. monadInv H0.
destruct (zlt (list_length_z x) Int.max_unsigned); inv EQ0.
+ rewrite transl_code'_transl_code in EQ.
exists x; exists true; split; auto. unfold fn_code. repeat constructor.
- exact transf_function_no_overflow.
Qed.
@@ -436,7 +437,7 @@ Lemma exec_straight_steps:
exists rs2,
exec_straight tge tf c rs1 m1' k rs2 m2'
/\ agree ms2 sp rs2
- /\ (edx_preserved ep i = true -> rs2#EDX = parent_sp s)) ->
+ /\ (it1_is_parent ep i = true -> rs2#EDX = parent_sp s)) ->
exists st',
plus step tge (State rs1 m1') E0 st' /\
match_states (Mach.State s fb sp c ms2 m2) st'.
@@ -455,7 +456,7 @@ Lemma exec_straight_steps_goto:
Genv.find_funct_ptr ge fb = Some (Internal f) ->
Mach.find_label lbl f.(Mach.fn_code) = Some c' ->
transl_code_at_pc ge (rs1 PC) fb f (i :: c) ep tf tc ->
- edx_preserved ep i = false ->
+ it1_is_parent ep i = false ->
(forall k c (TR: transl_instr f i ep k = OK c),
exists jmp, exists k', exists rs2,
exec_straight tge tf c rs1 m1' (jmp :: k') rs2 m2'
@@ -868,7 +869,7 @@ Opaque loadind.
- (* internal function *)
exploit functions_translated; eauto. intros [tf [A B]]. monadInv B.
- generalize EQ; intros EQ'. monadInv EQ'.
+ generalize EQ; intros EQ'. monadInv EQ'. rewrite transl_code'_transl_code in EQ0.
destruct (zlt (list_length_z x0) Int.max_unsigned); inversion EQ1. clear EQ1.
unfold store_stack in *.
exploit Mem.alloc_extends. eauto. eauto. apply Zle_refl. apply Zle_refl.
diff --git a/ia32/PrintOp.ml b/ia32/PrintOp.ml
index f76416a..7ddf42f 100644
--- a/ia32/PrintOp.ml
+++ b/ia32/PrintOp.ml
@@ -64,6 +64,7 @@ let print_operation reg pp = function
| Omove, [r1] -> reg pp r1
| Ointconst n, [] -> fprintf pp "%ld" (camlint_of_coqint n)
| Ofloatconst n, [] -> fprintf pp "%F" (camlfloat_of_coqfloat n)
+ | Oindirectsymbol id, [] -> fprintf pp "&%s" (extern_atom id)
| Ocast8signed, [r1] -> fprintf pp "int8signed(%a)" reg r1
| Ocast8unsigned, [r1] -> fprintf pp "int8unsigned(%a)" reg r1
| Ocast16signed, [r1] -> fprintf pp "int16signed(%a)" reg r1