From 54f97d1988f623ba7422e13a504caeb5701ba93c Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Fri, 21 Aug 2015 11:05:36 +0200 Subject: Refactoring of builtins and annotations in the back-end. Before, the back-end languages had distinct instructions - Iannot for annotations, taking structured expressions (annot_arg) as arguments, and producing no results' - Ibuiltin for other builtins, using simple pseudoregs/locations/registers as arguments and results. This branch enriches Ibuiltin instructions so that they take structured expressions (builtin_arg and builtin_res) as arguments and results. This way, - Annotations fit the general pattern of builtin functions, so Iannot instructions are removed. - EF_vload_global and EF_vstore_global become useless, as the same optimization can be achieved by EF_vload/vstore taking a structured argument of the "address of global" kind. - Better code can be generated for builtin_memcpy between stack locations, or volatile accesses to stack locations. Finally, this commit also introduces a new kind of external function, EF_debug, which is like EF_annot but produces no observable events. It will be used later to transport debug info through the back-end, without preventing optimizations. --- backend/Inlining.v | 21 ++++++++++++--------- 1 file changed, 12 insertions(+), 9 deletions(-) (limited to 'backend/Inlining.v') diff --git a/backend/Inlining.v b/backend/Inlining.v index 4f17d59c..98436bf5 100644 --- a/backend/Inlining.v +++ b/backend/Inlining.v @@ -203,15 +203,21 @@ Definition sop (ctx: context) (op: operation) := Definition saddr (ctx: context) (addr: addressing) := shift_stack_addressing (Int.repr ctx.(dstk)) addr. -Fixpoint sannotarg (ctx: context) (a: annot_arg reg) : annot_arg reg := +Fixpoint sbuiltinarg (ctx: context) (a: builtin_arg reg) : builtin_arg reg := match a with - | AA_base x => AA_base (sreg ctx x) - | AA_loadstack chunk ofs => AA_loadstack chunk (Int.add ofs (Int.repr ctx.(dstk))) - | AA_addrstack ofs => AA_addrstack (Int.add ofs (Int.repr ctx.(dstk))) - | AA_longofwords hi lo => AA_longofwords (sannotarg ctx hi) (sannotarg ctx lo) + | BA x => BA (sreg ctx x) + | BA_loadstack chunk ofs => BA_loadstack chunk (Int.add ofs (Int.repr ctx.(dstk))) + | BA_addrstack ofs => BA_addrstack (Int.add ofs (Int.repr ctx.(dstk))) + | BA_longofwords hi lo => BA_longofwords (sbuiltinarg ctx hi) (sbuiltinarg ctx lo) | _ => a end. +Definition sbuiltinres (ctx: context) (a: builtin_res reg) : builtin_res reg := + match a with + | BR x => BR (sreg ctx x) + | _ => BR_none + end. + (** The initial context, used to copy the CFG of a toplevel function. *) Definition initcontext (dpc dreg nreg: positive) (sz: Z) := @@ -390,10 +396,7 @@ Definition expand_instr (ctx: context) (pc: node) (i: instruction): mon unit := end | Ibuiltin ef args res s => set_instr (spc ctx pc) - (Ibuiltin ef (sregs ctx args) (sreg ctx res) (spc ctx s)) - | Iannot ef args s => - set_instr (spc ctx pc) - (Iannot ef (map (sannotarg ctx) args) (spc ctx s)) + (Ibuiltin ef (map (sbuiltinarg ctx) args) (sbuiltinres ctx res) (spc ctx s)) | Icond cond args s1 s2 => set_instr (spc ctx pc) (Icond cond (sregs ctx args) (spc ctx s1) (spc ctx s2)) -- cgit