From 9c7c84cc40eaacc1e2c13091165785cddecba5ad Mon Sep 17 00:00:00 2001 From: xleroy Date: Tue, 29 Jun 2010 08:27:14 +0000 Subject: Support for inlined built-ins. AST: add ef_inline flag to external functions. Selection: recognize calls to inlined built-ins and inline them as Sbuiltin. CminorSel to Asm: added Sbuiltin/Ibuiltin instruction. PrintAsm: adapted expansion of builtins. C2Clight: adapted detection of builtins. Conventions: refactored in a machine-independent part (backend/Conventions) and a machine-dependent part (ARCH/SYS/Conventions1). git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@1356 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e --- backend/Tailcallproof.v | 15 +++++++++++++-- 1 file changed, 13 insertions(+), 2 deletions(-) (limited to 'backend/Tailcallproof.v') diff --git a/backend/Tailcallproof.v b/backend/Tailcallproof.v index 2eed6e8d..11e6be20 100644 --- a/backend/Tailcallproof.v +++ b/backend/Tailcallproof.v @@ -24,7 +24,7 @@ Require Import Globalenvs. Require Import Smallstep. Require Import Registers. Require Import RTL. -Require Conventions. +Require Import Conventions. Require Import Tailcall. (** * Syntactic properties of the code transformation *) @@ -168,7 +168,7 @@ Lemma transf_instr_charact: transf_instr_spec f instr (transf_instr f pc instr). Proof. intros. unfold transf_instr. destruct instr; try constructor. - caseEq (is_return niter f n r && Conventions.tailcall_is_possible s && + caseEq (is_return niter f n r && tailcall_is_possible s && opt_typ_eq (sig_res s) (sig_res (fn_sig f))); intros. destruct (andb_prop _ _ H0). destruct (andb_prop _ _ H1). eapply transf_instr_tailcall; eauto. @@ -497,6 +497,17 @@ Proof. rewrite stacksize_preserved; auto. constructor. auto. apply regset_get_list; auto. auto. +(* builtin *) + TransfInstr. + assert (Val.lessdef_list (rs##args) (rs'##args)). apply regset_get_list; auto. + exploit external_call_mem_extends; eauto. + intros [v' [m'1 [A [B [C D]]]]]. + left. exists (State s' (transf_function f) (Vptr sp0 Int.zero) pc' (rs'#res <- v') m'1); split. + eapply exec_Ibuiltin; eauto. + eapply external_call_symbols_preserved; eauto. + exact symbols_preserved. exact varinfo_preserved. + econstructor; eauto. apply regset_set; auto. + (* cond true *) TransfInstr. left. exists (State s' (transf_function f) (Vptr sp0 Int.zero) ifso rs' m'); split. -- cgit