From 76ea1108be6f8b4ba9dc0118a13f685bcb62bc2b Mon Sep 17 00:00:00 2001 From: xleroy Date: Sat, 14 Sep 2013 16:24:30 +0000 Subject: Floats.v, Nan.v: hard-wire the general shape of binop_pl, so that no axioms are necessary, only two parameters (default_pl and choose_binop_pl). SelectDiv: optimize FP division by a power of 2. ConstpropOp: optimize 2.0 * x and x * 2.0 into x + x. git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@2326 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e --- backend/SelectDiv.vp | 18 +++++++++++++++++- 1 file changed, 17 insertions(+), 1 deletion(-) (limited to 'backend/SelectDiv.vp') diff --git a/backend/SelectDiv.vp b/backend/SelectDiv.vp index e60a97d..1d9168c 100644 --- a/backend/SelectDiv.vp +++ b/backend/SelectDiv.vp @@ -10,7 +10,7 @@ (* *) (* *********************************************************************) -(** Instruction selection for integer division and modulus *) +(** Instruction selection for division and modulus *) Require Import Coqlib. Require Import AST. @@ -154,3 +154,19 @@ Nondetfunction mods (e1: expr) (e2: expr) := | _ => mods_base e1 e2 end. +(** Floating-point division by a constant can also be turned into a FP + multiplication by the inverse constant, but only for powers of 2. *) + +Definition divfimm (e: expr) (n: float) := + match Float.exact_inverse n with + | Some n' => Eop Omulf (e ::: Eop (Ofloatconst n') Enil ::: Enil) + | None => Eop Odivf (e ::: Eop (Ofloatconst n) Enil ::: Enil) + end. + +Nondetfunction divf (e1: expr) (e2: expr) := + match e2 with + | Eop (Ofloatconst n2) Enil => divfimm e1 n2 + | _ => Eop Odivf (e1 ::: e2 ::: Enil) + end. + + -- cgit v1.2.3