Library Interval.Poly.Bound

This file is part of the CoqApprox formalization of rigorous polynomial approximation in Coq: http://tamadi.gforge.inria.fr/CoqApprox/
Copyright (C) 2010-2012, ENS de Lyon. Copyright (C) 2010-2016, Inria. Copyright (C) 2014-2016, IRIT.
This library is governed by the CeCILL-C license under French law and abiding by the rules of distribution of free software. You can use, modify and/or redistribute the library under the terms of the CeCILL-C license as circulated by CEA, CNRS and Inria at the following URL: http://www.cecill.info/
As a counterpart to the access to the source code and rights to copy, modify and redistribute granted by the license, users are provided only with a limited warranty and the library's author, the holder of the economic rights, and the successive licensors have only limited liability. See the COPYING file for more details.

From Coq Require Import ZArith Reals Psatz.
From mathcomp.ssreflect Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq fintype bigop.

Require Import Interval.
Require Import MathComp.
Require Import Datatypes.

Set Implicit Arguments.

The interface

Module Type PolyBound (I : IntervalOps) (Pol : PolyIntOps I).
Import Pol Pol.Notations.
Local Open Scope ipoly_scope.

Module J := IntervalExt I.

Parameter ComputeBound : Pol.U Pol.T I.type I.type.

Parameter ComputeBound_correct :
   u pi p,
  pi >:: p
  J.extension (PolR.horner tt p) (ComputeBound u pi).

Parameter ComputeBound_propagate :
   u pi,
  J.propagate (ComputeBound u pi).

End PolyBound.

Module PolyBoundThm
  (I : IntervalOps)
  (Pol : PolyIntOps I)
  (Bnd : PolyBound I Pol).

Import Pol.Notations Bnd.
Local Open Scope ipoly_scope.

Theorem ComputeBound_nth0 prec pi p X :
  pi >:: p
  X >: 0
   r : R, (Pol.nth pi 0) >: r
                ComputeBound prec pi X >: r.

End PolyBoundThm.

Naive implementation: Horner evaluation

Module PolyBoundHorner (I : IntervalOps) (Pol : PolyIntOps I)
  <: PolyBound I Pol.

Import Pol.Notations.
Local Open Scope ipoly_scope.
Module J := IntervalExt I.

Definition ComputeBound : Pol.U Pol.T I.type I.type :=
  Pol.horner.

Theorem ComputeBound_correct :
   prec pi p,
  pi >:: p
  J.extension (PolR.horner tt p) (ComputeBound prec pi).

Lemma ComputeBound_propagate :
   prec pi,
  J.propagate (ComputeBound prec pi).

End PolyBoundHorner.