Built with Alectryon, running Coq+SerAPI v8.10.0+0.7.0. Coq sources are in this panel; goals and messages will appear in the other. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑ Ctrl+↓ to navigate, Ctrl+🖱️ to focus.
(************************************************************************)
(*         *   The Coq Proof Assistant / The Coq Development Team       *)
(*  v      *   INRIA, CNRS and contributors - Copyright 1999-2018       *)
(* <O___,, *       (see CREDITS file for the list of authors)           *)
(*   \VV/  **************************************************************)
(*    //   *    This file is distributed under the terms of the         *)
(*         *     GNU Lesser General Public License Version 2.1          *)
(*         *     (see LICENSE file for the text of the license)         *)
(************************************************************************)

Require Import Rbase.
Require Import Rfunctions.
Require Import Ranalysis_reg.
Require Import Classical_Prop.
Local Open Scope R_scope.

Set Implicit Arguments.

(*****************************************************)

Each bounded subset of N has a maximal element

(*****************************************************)

Definition Nbound (I:nat -> Prop) : Prop :=
  exists n : nat, (forall i:nat, I i -> (i <= n)%nat).


forall z : Z, (0 <= z)%Z -> {n : nat | z = Z.of_nat n}

forall z : Z, (0 <= z)%Z -> {n : nat | z = Z.of_nat n}
intros; apply Z_of_nat_complete_inf; assumption. Qed.

forall I : nat -> Prop, (exists n : nat, I n) -> Nbound I -> {n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}

forall I : nat -> Prop, (exists n : nat, I n) -> Nbound I -> {n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x : R => exists i : nat, I i /\ INR i = x:R -> Prop

bound E
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x : R => exists i : nat, I i /\ INR i = x:R -> Prop
H1:bound E
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x : R => exists i : nat, I i /\ INR i = x:R -> Prop
H1:bound E

{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x : R => exists i : nat, I i /\ INR i = x:R -> Prop
H1:bound E

exists x : R, E x
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x : R => exists i : nat, I i /\ INR i = x:R -> Prop
H1:bound E
H2:exists x : R, E x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x : R => exists i : nat, I i /\ INR i = x:R -> Prop
H1:bound E
H2:exists x : R, E x

{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b

0 <= x
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
x0:R
H6:E x0
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b

0 <= x
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
x0:R
H6:E x0
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H7:E x0
HeqH7:H7 = H6

0 <= x
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
x0:R
H6:E x0
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
x1:nat
H8:I x1
H9:INR x1 = x0
HeqH7:ex_intro (fun i : nat => I i /\ INR i = x0) x1 (conj H8 H9) = H6

0 <= x
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x

{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1

x <= IZR (up x) - 1
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
x0:R
H9:E x0

x0 <= IZR (up x) - 1
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
x1:nat
H10:INR x1 <= x
H9:exists i : nat, I i /\ INR i = INR x1
H12:I x1

INR x1 <= IZR (up x) - 1
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
x1:nat
H10:INR x1 <= x
H9:exists i : nat, I i /\ INR i = INR x1
H12:I x1

INR (S x1) <= IZR (up x)
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
x1:nat
H10:INR x1 <= x
H9:exists i : nat, I i /\ INR i = INR x1
H12:I x1

(0 <= up x)%Z
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
x1:nat
H10:INR x1 <= x
H9:exists i : nat, I i /\ INR i = INR x1
H12:I x1
H14:(0 <= up x)%Z
INR (S x1) <= IZR (up x)
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
x1:nat
H10:INR x1 <= x
H9:exists i : nat, I i /\ INR i = INR x1
H12:I x1
H14:(0 <= up x)%Z

INR (S x1) <= IZR (up x)
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
x1:nat
H10:INR x1 <= x
H9:exists i : nat, I i /\ INR i = INR x1
H12:I x1
H14:(0 <= up x)%Z
x2:nat
H15:up x = Z.of_nat x2

INR (S x1) <= IZR (up x)
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
x1:nat
H10:INR x1 <= x
H9:exists i : nat, I i /\ INR i = INR x1
H12:I x1
H14:(0 <= up x)%Z
x2:nat
H15:up x = Z.of_nat x2

(x1 < x2)%nat
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1

{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1

x = IZR (up x) - 1
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1

{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1

(0 <= up x)%Z
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z

{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0

E x
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:E x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x

E x
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:E x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x

(forall y : R, E y -> y <= x - 1) -> E x
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
forall y : R, E y -> y <= x - 1
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:E x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
H13:forall y : R, E y -> y <= x - 1
H14:x <= x - 1

x - 1 < x -> E x
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
H13:forall y : R, E y -> y <= x - 1
H14:x <= x - 1
x - 1 < x
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
forall y : R, E y -> y <= x - 1
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:E x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
H13:forall y : R, E y -> y <= x - 1
H14:x <= x - 1

x - 1 < x
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
forall y : R, E y -> y <= x - 1
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:E x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x

forall y : R, E y -> y <= x - 1
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:E x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y < x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y

1 + y <= 1 + (x - 1)
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y = x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y
1 + y <= 1 + (x - 1)
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:E x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y < x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y

INR (S x1) <= x
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y = x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y
1 + y <= 1 + (x - 1)
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:E x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y < x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y

x = INR (Init.Nat.pred x0) -> INR (S x1) <= x
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y < x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y
x = INR (Init.Nat.pred x0)
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y = x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y
1 + y <= 1 + (x - 1)
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:E x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y < x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y

x = INR (Init.Nat.pred x0)
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y = x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y
1 + y <= 1 + (x - 1)
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:E x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y < x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y

INR (x0 - 1) = INR (Init.Nat.pred x0)
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y < x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y
(1 <= x0)%nat
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y = x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y
1 + y <= 1 + (x - 1)
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:E x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y < x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y

(1 <= x0)%nat
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y = x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y
1 + y <= 1 + (x - 1)
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:E x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
H8:up x = Z.of_nat 0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y < x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y

(1 <= 0)%nat
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat (S x0)
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y < x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y
Hrecx0:up x = Z.of_nat x0 -> (1 <= x0)%nat
(1 <= S x0)%nat
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y = x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y
1 + y <= 1 + (x - 1)
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:E x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:IZR (Z.of_nat 0) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
H8:up x = Z.of_nat 0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y < x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y

(1 <= 0)%nat
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat (S x0)
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y < x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y
Hrecx0:up x = Z.of_nat x0 -> (1 <= x0)%nat
(1 <= S x0)%nat
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y = x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y
1 + y <= 1 + (x - 1)
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:E x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x0 : R => exists i : nat, I i /\ INR i = x0:R -> Prop
H1:bound E
H2:exists x0 : R, E x0
x:R
H4:forall x0 : R, E x0 -> x0 <= x
H5:forall b : R, (forall x0 : R, E x0 -> x0 <= b) -> x <= b
H6:0 <= x
H3:0 > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
H8:up x = Z.of_nat 0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y < x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y

(1 <= 0)%nat
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat (S x0)
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y < x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y
Hrecx0:up x = Z.of_nat x0 -> (1 <= x0)%nat
(1 <= S x0)%nat
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y = x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y
1 + y <= 1 + (x - 1)
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:E x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat (S x0)
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y < x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y
Hrecx0:up x = Z.of_nat x0 -> (1 <= x0)%nat

(1 <= S x0)%nat
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y = x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y
1 + y <= 1 + (x - 1)
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:E x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H12:~ E x
y:R
H13:exists i : nat, I i /\ INR i = y
H14:y <= x
H15:y = x
x1:nat
H16:I x1 /\ INR x1 = y
H17:I x1
H18:INR x1 = y

1 + y <= 1 + (x - 1)
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:E x
{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x1 : R => exists i : nat, I i /\ INR i = x1:R -> Prop
H1:bound E
H2:exists x1 : R, E x1
x:R
H4:forall x1 : R, E x1 -> x1 <= x
H5:forall b : R, (forall x1 : R, E x1 -> x1 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:E x

{n : nat | I n /\ (forall i : nat, I i -> (i <= n)%nat)}
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:exists i : nat, I i /\ INR i = x
x1:nat
H12:I x1 /\ INR x1 = x
H14:I x1
H15:INR x1 = INR x0 - 1

INR x0 = INR x1 + 1
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:exists i : nat, I i /\ INR i = x
x1:nat
H12:I x1 /\ INR x1 = x
H14:I x1
H15:INR x1 = INR x0 - 1
H16:INR x0 = INR x1 + 1
I (Init.Nat.pred x0) /\ (forall i : nat, I i -> (i <= Init.Nat.pred x0)%nat)
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:exists i : nat, I i /\ INR i = x
x1:nat
H12:I x1 /\ INR x1 = x
H14:I x1
H15:INR x1 = INR x0 - 1
H16:INR x0 = INR x1 + 1

I (Init.Nat.pred x0) /\ (forall i : nat, I i -> (i <= Init.Nat.pred x0)%nat)
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:exists i : nat, I i /\ INR i = x
x1:nat
H12:I x1 /\ INR x1 = x
H14:I x1
H15:INR x1 = INR x0 - 1
H16:INR x0 = INR (S x1)
H17:x0 = S x1

I x1
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:exists i : nat, I i /\ INR i = x
x1:nat
H12:I x1 /\ INR x1 = x
H14:I x1
H15:INR x1 = INR x0 - 1
H16:INR x0 = INR (S x1)
H17:x0 = S x1
forall i : nat, I i -> (i <= x1)%nat
I:nat -> Prop
H:exists n : nat, I n
H0:Nbound I
E:=fun x2 : R => exists i : nat, I i /\ INR i = x2:R -> Prop
H1:bound E
H2:exists x2 : R, E x2
x:R
H4:forall x2 : R, E x2 -> x2 <= x
H5:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x <= b
H6:0 <= x
H3:IZR (up x) > x
H7:IZR (up x) - x <= 1
H9:x <= IZR (up x) - 1
H10:x = IZR (up x) - 1
H11:(0 <= up x)%Z
x0:nat
H8:up x = Z.of_nat x0
H13:exists i : nat, I i /\ INR i = x
x1:nat
H12:I x1 /\ INR x1 = x
H14:I x1
H15:INR x1 = INR x0 - 1
H16:INR x0 = INR (S x1)
H17:x0 = S x1

forall i : nat, I i -> (i <= x1)%nat
intros; apply INR_le; rewrite H15; rewrite <- H15; elim H12; intros; rewrite H20; apply H4; unfold E; exists i; split; [ assumption | reflexivity ]. Qed. (*******************************************)

Step functions

(*******************************************)

Definition open_interval (a b x:R) : Prop := a < x < b.
Definition co_interval (a b x:R) : Prop := a <= x < b.

Definition adapted_couple (f:R -> R) (a b:R) (l lf:Rlist) : Prop :=
  ordered_Rlist l /\
  pos_Rl l 0 = Rmin a b /\
  pos_Rl l (pred (Rlength l)) = Rmax a b /\
  Rlength l = S (Rlength lf) /\
  (forall i:nat,
    (i < pred (Rlength l))%nat ->
    constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i)))
    (pos_Rl lf i)).

Definition adapted_couple_opt (f:R -> R) (a b:R) (l lf:Rlist) :=
  adapted_couple f a b l lf /\
  (forall i:nat,
    (i < pred (Rlength lf))%nat ->
    pos_Rl lf i <> pos_Rl lf (S i) \/ f (pos_Rl l (S i)) <> pos_Rl lf i) /\
  (forall i:nat, (i < pred (Rlength l))%nat -> pos_Rl l i <> pos_Rl l (S i)).

Definition is_subdivision (f:R -> R) (a b:R) (l:Rlist) : Type :=
  { l0:Rlist & adapted_couple f a b l l0 }.

Definition IsStepFun (f:R -> R) (a b:R) : Type :=
  { l:Rlist & is_subdivision f a b l }.

Class of step functions

Record StepFun (a b:R) : Type := mkStepFun
  {fe :> R -> R; pre : IsStepFun fe a b}.

Definition subdivision (a b:R) (f:StepFun a b) : Rlist := projT1 (pre f).

Definition subdivision_val (a b:R) (f:StepFun a b) : Rlist :=
  match projT2 (pre f) with
    | existT _ a b => a
  end.

Fixpoint Int_SF (l k:Rlist) : R :=
  match l with
    | nil => 0
    | cons a l' =>
      match k with
        | nil => 0
        | cons x nil => 0
        | cons x (cons y k') => a * (y - x) + Int_SF l' (cons y k')
      end
  end.

Integral of step functions

Definition RiemannInt_SF (a b:R) (f:StepFun a b) : R :=
  match Rle_dec a b with
    | left _ => Int_SF (subdivision_val f) (subdivision f)
    | right _ => - Int_SF (subdivision_val f) (subdivision f)
  end.

(************************************)

Properties of step functions

(************************************)


forall (a b : R) (f : StepFun a b), adapted_couple f a b (subdivision f) (subdivision_val f)

forall (a b : R) (f : StepFun a b), adapted_couple f a b (subdivision f) (subdivision_val f)
intros a b f; unfold subdivision_val; case (projT2 (pre f)) as (x,H); apply H. Qed.

forall (a b : R) (f : R -> R) (l lf : Rlist), adapted_couple f a b l lf -> adapted_couple f b a l lf

forall (a b : R) (f : R -> R) (l lf : Rlist), adapted_couple f a b l lf -> adapted_couple f b a l lf
a, b:R
f:R -> R
l, lf:Rlist
H0:ordered_Rlist l
H2:pos_Rl l 0 = Rmin a b
H1:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H3:Rlength l = S (Rlength lf)
H5:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)

pos_Rl l 0 = Rmin b a
a, b:R
f:R -> R
l, lf:Rlist
H0:ordered_Rlist l
H2:pos_Rl l 0 = Rmin a b
H1:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H3:Rlength l = S (Rlength lf)
H5:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)
pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax b a
a, b:R
f:R -> R
l, lf:Rlist
H0:ordered_Rlist l
H2:pos_Rl l 0 = Rmin a b
H1:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H3:Rlength l = S (Rlength lf)
H5:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)
r:a <= b
r0:b <= a

a = b
a, b:R
f:R -> R
l, lf:Rlist
H0:ordered_Rlist l
H2:pos_Rl l 0 = Rmin a b
H1:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H3:Rlength l = S (Rlength lf)
H5:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)
n:~ a <= b
n0:~ b <= a
b = a
a, b:R
f:R -> R
l, lf:Rlist
H0:ordered_Rlist l
H2:pos_Rl l 0 = Rmin a b
H1:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H3:Rlength l = S (Rlength lf)
H5:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)
pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax b a
a, b:R
f:R -> R
l, lf:Rlist
H0:ordered_Rlist l
H2:pos_Rl l 0 = Rmin a b
H1:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H3:Rlength l = S (Rlength lf)
H5:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)
n:~ a <= b
n0:~ b <= a

b = a
a, b:R
f:R -> R
l, lf:Rlist
H0:ordered_Rlist l
H2:pos_Rl l 0 = Rmin a b
H1:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H3:Rlength l = S (Rlength lf)
H5:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)
pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax b a
a, b:R
f:R -> R
l, lf:Rlist
H0:ordered_Rlist l
H2:pos_Rl l 0 = Rmin a b
H1:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H3:Rlength l = S (Rlength lf)
H5:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)

pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax b a
a, b:R
f:R -> R
l, lf:Rlist
H0:ordered_Rlist l
H2:pos_Rl l 0 = Rmin a b
H1:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H3:Rlength l = S (Rlength lf)
H5:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)
r:a <= b
r0:b <= a

b = a
a, b:R
f:R -> R
l, lf:Rlist
H0:ordered_Rlist l
H2:pos_Rl l 0 = Rmin a b
H1:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H3:Rlength l = S (Rlength lf)
H5:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)
n:~ a <= b
n0:~ b <= a
a = b
a, b:R
f:R -> R
l, lf:Rlist
H0:ordered_Rlist l
H2:pos_Rl l 0 = Rmin a b
H1:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H3:Rlength l = S (Rlength lf)
H5:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)
n:~ a <= b
n0:~ b <= a

a = b
apply Rle_antisym; auto with real. Qed.

forall a b c : R, a <= b -> adapted_couple (fct_cte c) a b (cons a (cons b nil)) (cons c nil)

forall a b c : R, a <= b -> adapted_couple (fct_cte c) a b (cons a (cons b nil)) (cons c nil)
a, b, c:R
H:a <= b

ordered_Rlist (cons a (cons b nil))
a, b, c:R
H:a <= b
pos_Rl (cons a (cons b nil)) 0 = Rmin a b
a, b, c:R
H:a <= b
pos_Rl (cons a (cons b nil)) (Init.Nat.pred (Rlength (cons a (cons b nil)))) = Rmax a b
a, b, c:R
H:a <= b
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons b nil))))%nat -> constant_D_eq (fct_cte c) (open_interval (pos_Rl (cons a (cons b nil)) i) (pos_Rl (cons a (cons b nil)) (S i))) (pos_Rl (cons c nil) i)
a, b, c:R
H:a <= b

pos_Rl (cons a (cons b nil)) 0 = Rmin a b
a, b, c:R
H:a <= b
pos_Rl (cons a (cons b nil)) (Init.Nat.pred (Rlength (cons a (cons b nil)))) = Rmax a b
a, b, c:R
H:a <= b
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons b nil))))%nat -> constant_D_eq (fct_cte c) (open_interval (pos_Rl (cons a (cons b nil)) i) (pos_Rl (cons a (cons b nil)) (S i))) (pos_Rl (cons c nil) i)
a, b, c:R
H:a <= b

pos_Rl (cons a (cons b nil)) (Init.Nat.pred (Rlength (cons a (cons b nil)))) = Rmax a b
a, b, c:R
H:a <= b
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons b nil))))%nat -> constant_D_eq (fct_cte c) (open_interval (pos_Rl (cons a (cons b nil)) i) (pos_Rl (cons a (cons b nil)) (S i))) (pos_Rl (cons c nil) i)
a, b, c:R
H:a <= b

forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons b nil))))%nat -> constant_D_eq (fct_cte c) (open_interval (pos_Rl (cons a (cons b nil)) i) (pos_Rl (cons a (cons b nil)) (S i))) (pos_Rl (cons c nil) i)
unfold constant_D_eq, open_interval; intros; simpl in H0; inversion H0; [ reflexivity | elim (le_Sn_O _ H3) ]. Qed.

forall a b c : R, IsStepFun (fct_cte c) a b

forall a b c : R, IsStepFun (fct_cte c) a b
a, b, c:R
Hle:a <= b

{l : Rlist & is_subdivision (fct_cte c) a b l}
a, b, c:R
Hnle:~ a <= b
{l : Rlist & is_subdivision (fct_cte c) a b l}
a, b, c:R
Hnle:~ a <= b

{l : Rlist & is_subdivision (fct_cte c) a b l}
apply existT with (cons b (cons a nil)); unfold is_subdivision; apply existT with (cons c nil); apply StepFun_P2; apply StepFun_P3; auto with real. Qed.

forall (a b : R) (f : R -> R) (l : Rlist), is_subdivision f a b l -> is_subdivision f b a l

forall (a b : R) (f : R -> R) (l : Rlist), is_subdivision f a b l -> is_subdivision f b a l
a, b:R
f:R -> R
l, x:Rlist
H0:ordered_Rlist l
H1:pos_Rl l 0 = Rmin a b
H2:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H3:Rlength l = S (Rlength x)
H4:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl x i)

pos_Rl l 0 = Rmin b a
a, b:R
f:R -> R
l, x:Rlist
H0:ordered_Rlist l
H1:pos_Rl l 0 = Rmin a b
H2:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H3:Rlength l = S (Rlength x)
H4:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl x i)
pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax b a
a, b:R
f:R -> R
l, x:Rlist
H0:ordered_Rlist l
H1:pos_Rl l 0 = Rmin a b
H2:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H3:Rlength l = S (Rlength x)
H4:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl x i)

pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax b a
rewrite H2; apply Rmax_comm. Qed.

forall (f : R -> R) (a b : R), IsStepFun f a b -> IsStepFun f b a

forall (f : R -> R) (a b : R), IsStepFun f a b -> IsStepFun f b a
unfold IsStepFun; intros; elim X; intros; apply existT with x; apply StepFun_P5; assumption. Qed.

forall (a b r1 r2 r3 : R) (f : R -> R) (l lf : Rlist), a <= b -> adapted_couple f a b (cons r1 (cons r2 l)) (cons r3 lf) -> adapted_couple f r2 b (cons r2 l) lf

forall (a b r1 r2 r3 : R) (f : R -> R) (l lf : Rlist), a <= b -> adapted_couple f a b (cons r1 (cons r2 l)) (cons r3 lf) -> adapted_couple f r2 b (cons r2 l) lf
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)

Rmax a b = b
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b
ordered_Rlist (cons r2 l) /\ pos_Rl (cons r2 l) 0 = Rmin r2 b /\ pos_Rl (cons r2 l) (Init.Nat.pred (Rlength (cons r2 l))) = Rmax r2 b /\ Rlength (cons r2 l) = S (Rlength lf) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r2 l) i) (pos_Rl (cons r2 l) (S i))) (pos_Rl lf i))
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b

ordered_Rlist (cons r2 l) /\ pos_Rl (cons r2 l) 0 = Rmin r2 b /\ pos_Rl (cons r2 l) (Init.Nat.pred (Rlength (cons r2 l))) = Rmax r2 b /\ Rlength (cons r2 l) = S (Rlength lf) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r2 l) i) (pos_Rl (cons r2 l) (S i))) (pos_Rl lf i))
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b

r2 <= b
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b
H7:r2 <= b
ordered_Rlist (cons r2 l) /\ pos_Rl (cons r2 l) 0 = Rmin r2 b /\ pos_Rl (cons r2 l) (Init.Nat.pred (Rlength (cons r2 l))) = Rmax r2 b /\ Rlength (cons r2 l) = S (Rlength lf) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r2 l) i) (pos_Rl (cons r2 l) (S i))) (pos_Rl lf i))
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b
H7:r2 <= b

ordered_Rlist (cons r2 l) /\ pos_Rl (cons r2 l) 0 = Rmin r2 b /\ pos_Rl (cons r2 l) (Init.Nat.pred (Rlength (cons r2 l))) = Rmax r2 b /\ Rlength (cons r2 l) = S (Rlength lf) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r2 l) i) (pos_Rl (cons r2 l) (S i))) (pos_Rl lf i))
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b
H7:r2 <= b

ordered_Rlist (cons r2 l)
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b
H7:r2 <= b
pos_Rl (cons r2 l) 0 = Rmin r2 b
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b
H7:r2 <= b
pos_Rl (cons r2 l) (Init.Nat.pred (Rlength (cons r2 l))) = Rmax r2 b
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b
H7:r2 <= b
Rlength (cons r2 l) = S (Rlength lf)
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b
H7:r2 <= b
forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r2 l) i) (pos_Rl (cons r2 l) (S i))) (pos_Rl lf i)
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b
H7:r2 <= b

pos_Rl (cons r2 l) 0 = Rmin r2 b
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b
H7:r2 <= b
pos_Rl (cons r2 l) (Init.Nat.pred (Rlength (cons r2 l))) = Rmax r2 b
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b
H7:r2 <= b
Rlength (cons r2 l) = S (Rlength lf)
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b
H7:r2 <= b
forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r2 l) i) (pos_Rl (cons r2 l) (S i))) (pos_Rl lf i)
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b
H7:r2 <= b

pos_Rl (cons r2 l) (Init.Nat.pred (Rlength (cons r2 l))) = Rmax r2 b
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b
H7:r2 <= b
Rlength (cons r2 l) = S (Rlength lf)
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b
H7:r2 <= b
forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r2 l) i) (pos_Rl (cons r2 l) (S i))) (pos_Rl lf i)
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b
H7:r2 <= b

pos_Rl (cons r2 l) (Init.Nat.pred (Rlength (cons r2 l))) = b
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b
H7:r2 <= b
Rlength (cons r2 l) = S (Rlength lf)
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b
H7:r2 <= b
forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r2 l) i) (pos_Rl (cons r2 l) (S i))) (pos_Rl lf i)
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b
H7:r2 <= b

Rlength (cons r2 l) = S (Rlength lf)
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b
H7:r2 <= b
forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r2 l) i) (pos_Rl (cons r2 l) (S i))) (pos_Rl lf i)
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r1 (cons r2 l)) i) (pos_Rl (cons r1 (cons r2 l)) (S i))) (pos_Rl (cons r3 lf) i)
H5:Rmax a b = b
H7:r2 <= b

forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r2 l) i) (pos_Rl (cons r2 l) (S i))) (pos_Rl lf i)
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> forall x0 : R, pos_Rl (cons r1 (cons r2 l)) i0 < x0 < pos_Rl (cons r1 (cons r2 l)) (S i0) -> f x0 = pos_Rl (cons r3 lf) i0
H5:Rmax a b = b
H7:r2 <= b
i:nat
H0:(i < Init.Nat.pred (Rlength (cons r2 l)))%nat
x:R
H8:pos_Rl (cons r2 l) i < x < pos_Rl (cons r2 l) (S i)

(S i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> forall x0 : R, pos_Rl (cons r1 (cons r2 l)) i0 < x0 < pos_Rl (cons r1 (cons r2 l)) (S i0) -> f x0 = pos_Rl (cons r3 lf) i0
H5:Rmax a b = b
H7:r2 <= b
i:nat
H0:(i < Init.Nat.pred (Rlength (cons r2 l)))%nat
x:R
H8:pos_Rl (cons r2 l) i < x < pos_Rl (cons r2 l) (S i)
H9:(S i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat
f x = pos_Rl lf i
a, b, r1, r2, r3:R
f:R -> R
l, lf:Rlist
H:a <= b
H1:ordered_Rlist (cons r1 (cons r2 l))
H3:pos_Rl (cons r1 (cons r2 l)) 0 = Rmin a b
H2:pos_Rl (cons r1 (cons r2 l)) (Init.Nat.pred (Rlength (cons r1 (cons r2 l)))) = Rmax a b
H4:Rlength (cons r1 (cons r2 l)) = S (Rlength (cons r3 lf))
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat -> forall x0 : R, pos_Rl (cons r1 (cons r2 l)) i0 < x0 < pos_Rl (cons r1 (cons r2 l)) (S i0) -> f x0 = pos_Rl (cons r3 lf) i0
H5:Rmax a b = b
H7:r2 <= b
i:nat
H0:(i < Init.Nat.pred (Rlength (cons r2 l)))%nat
x:R
H8:pos_Rl (cons r2 l) i < x < pos_Rl (cons r2 l) (S i)
H9:(S i < Init.Nat.pred (Rlength (cons r1 (cons r2 l))))%nat

f x = pos_Rl lf i
assert (H10 := H6 _ H9); apply H10; assumption. Qed.

forall (f : R -> R) (l1 lf1 : Rlist) (a b : R), adapted_couple f a b l1 lf1 -> a = b -> Int_SF lf1 l1 = 0

forall (f : R -> R) (l1 lf1 : Rlist) (a b : R), adapted_couple f a b l1 lf1 -> a = b -> Int_SF lf1 l1 = 0
f:R -> R
l1:Rlist

forall (lf1 : Rlist) (a b : R), adapted_couple f a b nil lf1 -> a = b -> Int_SF lf1 nil = 0
f:R -> R
l1:Rlist
forall (r : R) (r0 : Rlist), (forall (lf1 : Rlist) (a b : R), adapted_couple f a b r0 lf1 -> a = b -> Int_SF lf1 r0 = 0) -> forall (lf1 : Rlist) (a b : R), adapted_couple f a b (cons r r0) lf1 -> a = b -> Int_SF lf1 (cons r r0) = 0
f:R -> R
l1:Rlist

forall (r : R) (r0 : Rlist), (forall (lf1 : Rlist) (a b : R), adapted_couple f a b r0 lf1 -> a = b -> Int_SF lf1 r0 = 0) -> forall (lf1 : Rlist) (a b : R), adapted_couple f a b (cons r r0) lf1 -> a = b -> Int_SF lf1 (cons r r0) = 0
f:R -> R
l1:Rlist
r:R
r0:Rlist

(forall (lf1 : Rlist) (a b : R), adapted_couple f a b nil lf1 -> a = b -> Int_SF lf1 nil = 0) -> forall (lf1 : Rlist) (a b : R), adapted_couple f a b (cons r nil) lf1 -> a = b -> Int_SF lf1 (cons r nil) = 0
f:R -> R
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b : R), adapted_couple f a b r2 lf1 -> a = b -> Int_SF lf1 r2 = 0) -> forall (lf1 : Rlist) (a b : R), adapted_couple f a b (cons r r2) lf1 -> a = b -> Int_SF lf1 (cons r r2) = 0) -> (forall (lf1 : Rlist) (a b : R), adapted_couple f a b (cons r1 r2) lf1 -> a = b -> Int_SF lf1 (cons r1 r2) = 0) -> forall (lf1 : Rlist) (a b : R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a = b -> Int_SF lf1 (cons r (cons r1 r2)) = 0
f:R -> R
l1:Rlist
r:R
r0:Rlist
H:forall (lf1 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 nil lf1 -> a0 = b0 -> Int_SF lf1 nil = 0
a, b:R
H0:adapted_couple f a b (cons r nil) nil
H1:a = b

Int_SF nil (cons r nil) = 0
f:R -> R
l1:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 nil lf0 -> a0 = b0 -> Int_SF lf0 nil = 0
r1:R
lf1:Rlist
a, b:R
H0:adapted_couple f a b (cons r nil) (cons r1 lf1)
H1:a = b
Hreclf1:adapted_couple f a b (cons r nil) lf1 -> Int_SF lf1 (cons r nil) = 0
Int_SF (cons r1 lf1) (cons r nil) = 0
f:R -> R
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b : R), adapted_couple f a b r2 lf1 -> a = b -> Int_SF lf1 r2 = 0) -> forall (lf1 : Rlist) (a b : R), adapted_couple f a b (cons r r2) lf1 -> a = b -> Int_SF lf1 (cons r r2) = 0) -> (forall (lf1 : Rlist) (a b : R), adapted_couple f a b (cons r1 r2) lf1 -> a = b -> Int_SF lf1 (cons r1 r2) = 0) -> forall (lf1 : Rlist) (a b : R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a = b -> Int_SF lf1 (cons r (cons r1 r2)) = 0
f:R -> R
l1:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 nil lf0 -> a0 = b0 -> Int_SF lf0 nil = 0
r1:R
lf1:Rlist
a, b:R
H0:adapted_couple f a b (cons r nil) (cons r1 lf1)
H1:a = b
Hreclf1:adapted_couple f a b (cons r nil) lf1 -> Int_SF lf1 (cons r nil) = 0

Int_SF (cons r1 lf1) (cons r nil) = 0
f:R -> R
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b : R), adapted_couple f a b r2 lf1 -> a = b -> Int_SF lf1 r2 = 0) -> forall (lf1 : Rlist) (a b : R), adapted_couple f a b (cons r r2) lf1 -> a = b -> Int_SF lf1 (cons r r2) = 0) -> (forall (lf1 : Rlist) (a b : R), adapted_couple f a b (cons r1 r2) lf1 -> a = b -> Int_SF lf1 (cons r1 r2) = 0) -> forall (lf1 : Rlist) (a b : R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a = b -> Int_SF lf1 (cons r (cons r1 r2)) = 0
f:R -> R
l1:Rlist
r:R
r0:Rlist

forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b : R), adapted_couple f a b r2 lf1 -> a = b -> Int_SF lf1 r2 = 0) -> forall (lf1 : Rlist) (a b : R), adapted_couple f a b (cons r r2) lf1 -> a = b -> Int_SF lf1 (cons r r2) = 0) -> (forall (lf1 : Rlist) (a b : R), adapted_couple f a b (cons r1 r2) lf1 -> a = b -> Int_SF lf1 (cons r1 r2) = 0) -> forall (lf1 : Rlist) (a b : R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a = b -> Int_SF lf1 (cons r (cons r1 r2)) = 0
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H:(forall (lf1 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 r2 lf1 -> a0 = b0 -> Int_SF lf1 r2 = 0) -> forall (lf1 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r r2) lf1 -> a0 = b0 -> Int_SF lf1 (cons r r2) = 0
H0:forall (lf1 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r1 r2) lf1 -> a0 = b0 -> Int_SF lf1 (cons r1 r2) = 0
a, b:R
H1:adapted_couple f a b (cons r (cons r1 r2)) nil
H2:a = b

Int_SF nil (cons r (cons r1 r2)) = 0
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H:(forall (lf2 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 r2 lf2 -> a0 = b0 -> Int_SF lf2 r2 = 0) -> forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r r2) = 0
H0:forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r1 r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r1 r2) = 0
r3:R
lf1:Rlist
a, b:R
H1:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H2:a = b
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> Int_SF lf1 (cons r (cons r1 r2)) = 0
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = 0
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H:(forall (lf2 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 r2 lf2 -> a0 = b0 -> Int_SF lf2 r2 = 0) -> forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r r2) = 0
H0:forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r1 r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r1 r2) = 0
r3:R
lf1:Rlist
a, b:R
H1:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H2:a = b
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> Int_SF lf1 (cons r (cons r1 r2)) = 0

Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = 0
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H:(forall (lf2 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 r2 lf2 -> a0 = b0 -> Int_SF lf2 r2 = 0) -> forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r r2) = 0
H0:forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r1 r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r1 r2) = 0
r3:R
lf1:Rlist
a, b:R
H1:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H2:a = b
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> Int_SF lf1 (cons r (cons r1 r2)) = 0

r = r1 -> r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = 0
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H:(forall (lf2 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 r2 lf2 -> a0 = b0 -> Int_SF lf2 r2 = 0) -> forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r r2) = 0
H0:forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r1 r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r1 r2) = 0
r3:R
lf1:Rlist
a, b:R
H1:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H2:a = b
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> Int_SF lf1 (cons r (cons r1 r2)) = 0
r = r1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H:(forall (lf2 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 r2 lf2 -> a0 = b0 -> Int_SF lf2 r2 = 0) -> forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r r2) = 0
H0:forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r1 r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r1 r2) = 0
r3:R
lf1:Rlist
a, b:R
H1:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H2:a = b
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> Int_SF lf1 (cons r (cons r1 r2)) = 0
H3:r = r1

r3 * (r1 - r1) + 0 = 0
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H:(forall (lf2 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 r2 lf2 -> a0 = b0 -> Int_SF lf2 r2 = 0) -> forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r r2) = 0
H0:forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r1 r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r1 r2) = 0
r3:R
lf1:Rlist
a, b:R
H1:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H2:a = b
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> Int_SF lf1 (cons r (cons r1 r2)) = 0
H3:r = r1
adapted_couple f r b (cons r1 r2) lf1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H:(forall (lf2 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 r2 lf2 -> a0 = b0 -> Int_SF lf2 r2 = 0) -> forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r r2) = 0
H0:forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r1 r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r1 r2) = 0
r3:R
lf1:Rlist
a, b:R
H1:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H2:a = b
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> Int_SF lf1 (cons r (cons r1 r2)) = 0
H3:r = r1
r = b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H:(forall (lf2 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 r2 lf2 -> a0 = b0 -> Int_SF lf2 r2 = 0) -> forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r r2) = 0
H0:forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r1 r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r1 r2) = 0
r3:R
lf1:Rlist
a, b:R
H1:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H2:a = b
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> Int_SF lf1 (cons r (cons r1 r2)) = 0
r = r1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H:(forall (lf2 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 r2 lf2 -> a0 = b0 -> Int_SF lf2 r2 = 0) -> forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r r2) = 0
H0:forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r1 r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r1 r2) = 0
r3:R
lf1:Rlist
a, b:R
H1:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H2:a = b
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> Int_SF lf1 (cons r (cons r1 r2)) = 0
H3:r = r1

adapted_couple f r b (cons r1 r2) lf1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H:(forall (lf2 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 r2 lf2 -> a0 = b0 -> Int_SF lf2 r2 = 0) -> forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r r2) = 0
H0:forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r1 r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r1 r2) = 0
r3:R
lf1:Rlist
a, b:R
H1:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H2:a = b
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> Int_SF lf1 (cons r (cons r1 r2)) = 0
H3:r = r1
r = b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H:(forall (lf2 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 r2 lf2 -> a0 = b0 -> Int_SF lf2 r2 = 0) -> forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r r2) = 0
H0:forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r1 r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r1 r2) = 0
r3:R
lf1:Rlist
a, b:R
H1:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H2:a = b
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> Int_SF lf1 (cons r (cons r1 r2)) = 0
r = r1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H:(forall (lf2 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 r2 lf2 -> a0 = b0 -> Int_SF lf2 r2 = 0) -> forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r r2) = 0
H0:forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r1 r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r1 r2) = 0
r3:R
lf1:Rlist
a, b:R
H1:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H2:a = b
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> Int_SF lf1 (cons r (cons r1 r2)) = 0
H3:r = r1

r = b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H:(forall (lf2 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 r2 lf2 -> a0 = b0 -> Int_SF lf2 r2 = 0) -> forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r r2) = 0
H0:forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r1 r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r1 r2) = 0
r3:R
lf1:Rlist
a, b:R
H1:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H2:a = b
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> Int_SF lf1 (cons r (cons r1 r2)) = 0
r = r1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H:(forall (lf2 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 r2 lf2 -> a0 = b0 -> Int_SF lf2 r2 = 0) -> forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r r2) = 0
H0:forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r1 r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r1 r2) = 0
r3:R
lf1:Rlist
a, b:R
H1:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H2:a = b
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> Int_SF lf1 (cons r (cons r1 r2)) = 0

r = r1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H:(forall (lf2 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 r2 lf2 -> a0 = b0 -> Int_SF lf2 r2 = 0) -> forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r r2) = 0
H0:forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r1 r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r1 r2) = 0
r3:R
lf1:Rlist
a, b:R
H1:ordered_Rlist (cons r (cons r1 r2)) /\ pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b /\ pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b /\ Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1)) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i))
H2:a = b
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> Int_SF lf1 (cons r (cons r1 r2)) = 0
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)

r <= r1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H:(forall (lf2 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 r2 lf2 -> a0 = b0 -> Int_SF lf2 r2 = 0) -> forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r r2) = 0
H0:forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r1 r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r1 r2) = 0
r3:R
lf1:Rlist
a, b:R
H1:ordered_Rlist (cons r (cons r1 r2)) /\ pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b /\ pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b /\ Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1)) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i))
H2:a = b
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> Int_SF lf1 (cons r (cons r1 r2)) = 0
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
r1 <= r
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H:(forall (lf2 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 r2 lf2 -> a0 = b0 -> Int_SF lf2 r2 = 0) -> forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r r2) = 0
H0:forall (lf0 : Rlist) (a0 b0 : R), adapted_couple f a0 b0 (cons r1 r2) lf0 -> a0 = b0 -> Int_SF lf0 (cons r1 r2) = 0
r3:R
lf1:Rlist
a, b:R
H1:ordered_Rlist (cons r (cons r1 r2)) /\ pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b /\ pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b /\ Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1)) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i))
H2:a = b
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> Int_SF lf1 (cons r (cons r1 r2)) = 0
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)

r1 <= r
simpl in H5; rewrite H2 in H5; rewrite H5; replace (Rmin b b) with (Rmax a b); [ rewrite <- H4; apply RList_P7; [ assumption | simpl; right; left; reflexivity ] | unfold Rmin, Rmax; case (Rle_dec b b); case (Rle_dec a b); intros; try assumption || reflexivity ]. Qed.

forall (a b : R) (f : R -> R) (l lf : Rlist), adapted_couple f a b l lf -> a <> b -> (2 <= Rlength l)%nat

forall (a b : R) (f : R -> R) (l lf : Rlist), adapted_couple f a b l lf -> a <> b -> (2 <= Rlength l)%nat
intros; unfold adapted_couple in H; decompose [and] H; clear H; induction l as [| r l Hrecl]; [ simpl in H4; discriminate | induction l as [| r0 l Hrecl0]; [ simpl in H3; simpl in H2; generalize H3; generalize H2; unfold Rmin, Rmax; case (Rle_dec a b); intros; elim H0; rewrite <- H5; rewrite <- H7; reflexivity | simpl; do 2 apply le_n_S; apply le_O_n ] ]. Qed.

forall (f : R -> R) (l lf : Rlist) (a b : R), a <= b -> adapted_couple f a b l lf -> exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'

forall (f : R -> R) (l lf : Rlist) (a b : R), a <= b -> adapted_couple f a b l lf -> exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
f:R -> R
l:Rlist

forall (lf : Rlist) (a b : R), a <= b -> adapted_couple f a b nil lf -> exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
f:R -> R
l:Rlist
forall (r : R) (r0 : Rlist), (forall (lf : Rlist) (a b : R), a <= b -> adapted_couple f a b r0 lf -> exists l' lf' : Rlist, adapted_couple_opt f a b l' lf') -> forall (lf : Rlist) (a b : R), a <= b -> adapted_couple f a b (cons r r0) lf -> exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
f:R -> R
l:Rlist

forall (r : R) (r0 : Rlist), (forall (lf : Rlist) (a b : R), a <= b -> adapted_couple f a b r0 lf -> exists l' lf' : Rlist, adapted_couple_opt f a b l' lf') -> forall (lf : Rlist) (a b : R), a <= b -> adapted_couple f a b (cons r r0) lf -> exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) lf
H2:a = b

exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) lf
H2:a <> b
exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) lf
H2:a = b

pos_Rl (cons a nil) 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) lf
H2:a = b
pos_Rl (cons a nil) (Init.Nat.pred (Rlength (cons a nil))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) lf
H2:a <> b
exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) lf
H2:a = b

pos_Rl (cons a nil) (Init.Nat.pred (Rlength (cons a nil))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) lf
H2:a <> b
exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) lf
H2:a <> b

exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) nil
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)

exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
Hreclf:adapted_couple f a b (cons r r0) lf -> exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
Hreclf:adapted_couple f a b (cons r r0) lf -> exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'

exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)

adapted_couple f t2 b r0 lf
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf

exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf

t2 <= b -> exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
H6:exists l' lf' : Rlist, adapted_couple_opt f t2 b l' lf'
Hyp_eq:t1 = t2

exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
H6:exists l' lf' : Rlist, adapted_couple_opt f t2 b l' lf'
Hyp_eq:t1 <> t2
exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
H6:exists l' lf' : Rlist, adapted_couple_opt f t2 b l' lf'
Hyp_eq:t1 = t2

exists l' lf' : Rlist, adapted_couple_opt f t2 b l' lf'
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
H6:exists l' lf' : Rlist, adapted_couple_opt f t2 b l' lf'
Hyp_eq:t1 = t2
t2 = a
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
H6:exists l' lf' : Rlist, adapted_couple_opt f t2 b l' lf'
Hyp_eq:t1 <> t2
exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
H6:exists l' lf' : Rlist, adapted_couple_opt f t2 b l' lf'
Hyp_eq:t1 = t2

t2 = a
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
H6:exists l' lf' : Rlist, adapted_couple_opt f t2 b l' lf'
Hyp_eq:t1 <> t2
exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
H6:exists l' lf' : Rlist, adapted_couple_opt f t2 b l' lf'
Hyp_eq:t1 <> t2

exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b

exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 <> b
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b

ordered_Rlist (cons a (cons b nil))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
pos_Rl (cons a (cons b nil)) 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
pos_Rl (cons a (cons b nil)) (Init.Nat.pred (Rlength (cons a (cons b nil)))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons b nil))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons a (cons b nil)) i) (pos_Rl (cons a (cons b nil)) (S i))) (pos_Rl (cons r1 nil) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 nil)))%nat -> pos_Rl (cons r1 nil) i <> pos_Rl (cons r1 nil) (S i) \/ f (pos_Rl (cons a (cons b nil)) (S i)) <> pos_Rl (cons r1 nil) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons b nil))))%nat -> pos_Rl (cons a (cons b nil)) i <> pos_Rl (cons a (cons b nil)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 <> b
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b

pos_Rl (cons a (cons b nil)) 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
pos_Rl (cons a (cons b nil)) (Init.Nat.pred (Rlength (cons a (cons b nil)))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons b nil))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons a (cons b nil)) i) (pos_Rl (cons a (cons b nil)) (S i))) (pos_Rl (cons r1 nil) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 nil)))%nat -> pos_Rl (cons r1 nil) i <> pos_Rl (cons r1 nil) (S i) \/ f (pos_Rl (cons a (cons b nil)) (S i)) <> pos_Rl (cons r1 nil) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons b nil))))%nat -> pos_Rl (cons a (cons b nil)) i <> pos_Rl (cons a (cons b nil)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 <> b
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b

pos_Rl (cons a (cons b nil)) (Init.Nat.pred (Rlength (cons a (cons b nil)))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons b nil))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons a (cons b nil)) i) (pos_Rl (cons a (cons b nil)) (S i))) (pos_Rl (cons r1 nil) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 nil)))%nat -> pos_Rl (cons r1 nil) i <> pos_Rl (cons r1 nil) (S i) \/ f (pos_Rl (cons a (cons b nil)) (S i)) <> pos_Rl (cons r1 nil) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons b nil))))%nat -> pos_Rl (cons a (cons b nil)) i <> pos_Rl (cons a (cons b nil)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 <> b
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b

forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons b nil))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons a (cons b nil)) i) (pos_Rl (cons a (cons b nil)) (S i))) (pos_Rl (cons r1 nil) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 nil)))%nat -> pos_Rl (cons r1 nil) i <> pos_Rl (cons r1 nil) (S i) \/ f (pos_Rl (cons a (cons b nil)) (S i)) <> pos_Rl (cons r1 nil) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons b nil))))%nat -> pos_Rl (cons a (cons b nil)) i <> pos_Rl (cons a (cons b nil)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 <> b
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
i:nat
H8:(i < 1)%nat
H10:i = 0%nat

constant_D_eq f (open_interval (pos_Rl (cons a (cons b nil)) 0) (pos_Rl (cons a (cons b nil)) 1)) (pos_Rl (cons r1 nil) 0)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
i:nat
H8:(i < 1)%nat
m:nat
H10:(S i <= 0)%nat
H9:m = 0%nat
constant_D_eq f (open_interval (pos_Rl (cons a (cons b nil)) i) (pos_Rl (cons a (cons b nil)) (S i))) (pos_Rl (cons r1 nil) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 nil)))%nat -> pos_Rl (cons r1 nil) i <> pos_Rl (cons r1 nil) (S i) \/ f (pos_Rl (cons a (cons b nil)) (S i)) <> pos_Rl (cons r1 nil) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons b nil))))%nat -> pos_Rl (cons a (cons b nil)) i <> pos_Rl (cons a (cons b nil)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 <> b
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H1:ordered_Rlist (cons t1 (cons t2 t3)) /\ pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b /\ pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b /\ Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf)) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0))
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
i:nat
H8:(i < 1)%nat
H10:i = 0%nat
x:R
H9:a < x < b
H11:ordered_Rlist (cons t1 (cons t2 t3))
H13:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H12:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H14:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)

(0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H1:ordered_Rlist (cons t1 (cons t2 t3)) /\ pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b /\ pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b /\ Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf)) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0))
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
i:nat
H8:(i < 1)%nat
H10:i = 0%nat
x:R
H9:a < x < b
H11:ordered_Rlist (cons t1 (cons t2 t3))
H13:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H12:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H14:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
open_interval (pos_Rl (cons t1 (cons t2 t3)) 0) (pos_Rl (cons t1 (cons t2 t3)) 1) x
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
i:nat
H8:(i < 1)%nat
m:nat
H10:(S i <= 0)%nat
H9:m = 0%nat
constant_D_eq f (open_interval (pos_Rl (cons a (cons b nil)) i) (pos_Rl (cons a (cons b nil)) (S i))) (pos_Rl (cons r1 nil) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 nil)))%nat -> pos_Rl (cons r1 nil) i <> pos_Rl (cons r1 nil) (S i) \/ f (pos_Rl (cons a (cons b nil)) (S i)) <> pos_Rl (cons r1 nil) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons b nil))))%nat -> pos_Rl (cons a (cons b nil)) i <> pos_Rl (cons a (cons b nil)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 <> b
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H1:ordered_Rlist (cons t1 (cons t2 t3)) /\ pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b /\ pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b /\ Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf)) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0))
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
i:nat
H8:(i < 1)%nat
H10:i = 0%nat
x:R
H9:a < x < b
H11:ordered_Rlist (cons t1 (cons t2 t3))
H13:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H12:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H14:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)

open_interval (pos_Rl (cons t1 (cons t2 t3)) 0) (pos_Rl (cons t1 (cons t2 t3)) 1) x
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
i:nat
H8:(i < 1)%nat
m:nat
H10:(S i <= 0)%nat
H9:m = 0%nat
constant_D_eq f (open_interval (pos_Rl (cons a (cons b nil)) i) (pos_Rl (cons a (cons b nil)) (S i))) (pos_Rl (cons r1 nil) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 nil)))%nat -> pos_Rl (cons r1 nil) i <> pos_Rl (cons r1 nil) (S i) \/ f (pos_Rl (cons a (cons b nil)) (S i)) <> pos_Rl (cons r1 nil) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons b nil))))%nat -> pos_Rl (cons a (cons b nil)) i <> pos_Rl (cons a (cons b nil)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 <> b
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
i:nat
H8:(i < 1)%nat
m:nat
H10:(S i <= 0)%nat
H9:m = 0%nat

constant_D_eq f (open_interval (pos_Rl (cons a (cons b nil)) i) (pos_Rl (cons a (cons b nil)) (S i))) (pos_Rl (cons r1 nil) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 nil)))%nat -> pos_Rl (cons r1 nil) i <> pos_Rl (cons r1 nil) (S i) \/ f (pos_Rl (cons a (cons b nil)) (S i)) <> pos_Rl (cons r1 nil) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons b nil))))%nat -> pos_Rl (cons a (cons b nil)) i <> pos_Rl (cons a (cons b nil)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 <> b
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b

forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 nil)))%nat -> pos_Rl (cons r1 nil) i <> pos_Rl (cons r1 nil) (S i) \/ f (pos_Rl (cons a (cons b nil)) (S i)) <> pos_Rl (cons r1 nil) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons b nil))))%nat -> pos_Rl (cons a (cons b nil)) i <> pos_Rl (cons a (cons b nil)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 <> b
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 = b

forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons b nil))))%nat -> pos_Rl (cons a (cons b nil)) i <> pos_Rl (cons a (cons b nil)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 <> b
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 <> b

exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 <> b

Rmin t2 b = t2
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l', lf':Rlist
H6:adapted_couple_opt f t2 b l' lf'
H7:t2 <> b
Hyp_min:Rmin t2 b = t2

exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf' : Rlist, adapted_couple_opt f a0 b0 l'0 lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' nil
H8:(forall i : nat, (i < Init.Nat.pred (Rlength nil))%nat -> pos_Rl nil i <> pos_Rl nil (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl nil i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)

exists l'0 lf' : Rlist, adapted_couple_opt f a b l'0 lf'
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
Hreclf':adapted_couple f t2 b l' lf' -> (forall i : nat, (i < Init.Nat.pred (Rlength lf'))%nat -> pos_Rl lf' i <> pos_Rl lf' (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl lf' i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i)) -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
Hreclf':adapted_couple f t2 b l' lf' -> (forall i : nat, (i < Init.Nat.pred (Rlength lf'))%nat -> pos_Rl lf' i <> pos_Rl lf' (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl lf' i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i)) -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0

exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2

exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1

exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)

ordered_Rlist (cons t1 (cons s2 s3))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
pos_Rl (cons t1 (cons s2 s3)) 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
pos_Rl (cons t1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons t1 (cons s2 s3)))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
Rlength (cons t1 (cons s2 s3)) = S (Rlength (cons r1 lf'))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons s2 s3)) i) (pos_Rl (cons t1 (cons s2 s3)) (S i))) (pos_Rl (cons r1 lf') i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 lf')))%nat -> pos_Rl (cons r1 lf') i <> pos_Rl (cons r1 lf') (S i) \/ f (pos_Rl (cons t1 (cons s2 s3)) (S i)) <> pos_Rl (cons r1 lf') i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < S (Rlength s3))%nat

pos_Rl (cons t1 (cons s2 s3)) 0 <= pos_Rl (cons t1 (cons s2 s3)) 1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
i:nat
H1:(S i < S (Rlength s3))%nat
Hreci:(i < S (Rlength s3))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <= pos_Rl (cons t1 (cons s2 s3)) (S i)
pos_Rl (cons t1 (cons s2 s3)) (S i) <= pos_Rl (cons t1 (cons s2 s3)) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
pos_Rl (cons t1 (cons s2 s3)) 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
pos_Rl (cons t1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons t1 (cons s2 s3)))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
Rlength (cons t1 (cons s2 s3)) = S (Rlength (cons r1 lf'))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons s2 s3)) i) (pos_Rl (cons t1 (cons s2 s3)) (S i))) (pos_Rl (cons r1 lf') i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 lf')))%nat -> pos_Rl (cons r1 lf') i <> pos_Rl (cons r1 lf') (S i) \/ f (pos_Rl (cons t1 (cons s2 s3)) (S i)) <> pos_Rl (cons r1 lf') i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < S (Rlength s3))%nat

t1 <= s1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < S (Rlength s3))%nat
s1 <= s2
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
i:nat
H1:(S i < S (Rlength s3))%nat
Hreci:(i < S (Rlength s3))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <= pos_Rl (cons t1 (cons s2 s3)) (S i)
pos_Rl (cons t1 (cons s2 s3)) (S i) <= pos_Rl (cons t1 (cons s2 s3)) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
pos_Rl (cons t1 (cons s2 s3)) 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
pos_Rl (cons t1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons t1 (cons s2 s3)))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
Rlength (cons t1 (cons s2 s3)) = S (Rlength (cons r1 lf'))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons s2 s3)) i) (pos_Rl (cons t1 (cons s2 s3)) (S i))) (pos_Rl (cons r1 lf') i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 lf')))%nat -> pos_Rl (cons r1 lf') i <> pos_Rl (cons r1 lf') (S i) \/ f (pos_Rl (cons t1 (cons s2 s3)) (S i)) <> pos_Rl (cons r1 lf') i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < S (Rlength s3))%nat

t1 <= t2
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < S (Rlength s3))%nat
t2 = s1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < S (Rlength s3))%nat
s1 <= s2
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
i:nat
H1:(S i < S (Rlength s3))%nat
Hreci:(i < S (Rlength s3))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <= pos_Rl (cons t1 (cons s2 s3)) (S i)
pos_Rl (cons t1 (cons s2 s3)) (S i) <= pos_Rl (cons t1 (cons s2 s3)) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
pos_Rl (cons t1 (cons s2 s3)) 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
pos_Rl (cons t1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons t1 (cons s2 s3)))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
Rlength (cons t1 (cons s2 s3)) = S (Rlength (cons r1 lf'))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons s2 s3)) i) (pos_Rl (cons t1 (cons s2 s3)) (S i))) (pos_Rl (cons r1 lf') i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 lf')))%nat -> pos_Rl (cons r1 lf') i <> pos_Rl (cons r1 lf') (S i) \/ f (pos_Rl (cons t1 (cons s2 s3)) (S i)) <> pos_Rl (cons r1 lf') i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < S (Rlength s3))%nat

(0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < S (Rlength s3))%nat
t2 = s1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < S (Rlength s3))%nat
s1 <= s2
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
i:nat
H1:(S i < S (Rlength s3))%nat
Hreci:(i < S (Rlength s3))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <= pos_Rl (cons t1 (cons s2 s3)) (S i)
pos_Rl (cons t1 (cons s2 s3)) (S i) <= pos_Rl (cons t1 (cons s2 s3)) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
pos_Rl (cons t1 (cons s2 s3)) 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
pos_Rl (cons t1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons t1 (cons s2 s3)))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
Rlength (cons t1 (cons s2 s3)) = S (Rlength (cons r1 lf'))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons s2 s3)) i) (pos_Rl (cons t1 (cons s2 s3)) (S i))) (pos_Rl (cons r1 lf') i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 lf')))%nat -> pos_Rl (cons r1 lf') i <> pos_Rl (cons r1 lf') (S i) \/ f (pos_Rl (cons t1 (cons s2 s3)) (S i)) <> pos_Rl (cons r1 lf') i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < S (Rlength s3))%nat

t2 = s1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < S (Rlength s3))%nat
s1 <= s2
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
i:nat
H1:(S i < S (Rlength s3))%nat
Hreci:(i < S (Rlength s3))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <= pos_Rl (cons t1 (cons s2 s3)) (S i)
pos_Rl (cons t1 (cons s2 s3)) (S i) <= pos_Rl (cons t1 (cons s2 s3)) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
pos_Rl (cons t1 (cons s2 s3)) 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
pos_Rl (cons t1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons t1 (cons s2 s3)))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
Rlength (cons t1 (cons s2 s3)) = S (Rlength (cons r1 lf'))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons s2 s3)) i) (pos_Rl (cons t1 (cons s2 s3)) (S i))) (pos_Rl (cons r1 lf') i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 lf')))%nat -> pos_Rl (cons r1 lf') i <> pos_Rl (cons r1 lf') (S i) \/ f (pos_Rl (cons t1 (cons s2 s3)) (S i)) <> pos_Rl (cons r1 lf') i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < S (Rlength s3))%nat

s1 <= s2
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
i:nat
H1:(S i < S (Rlength s3))%nat
Hreci:(i < S (Rlength s3))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <= pos_Rl (cons t1 (cons s2 s3)) (S i)
pos_Rl (cons t1 (cons s2 s3)) (S i) <= pos_Rl (cons t1 (cons s2 s3)) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
pos_Rl (cons t1 (cons s2 s3)) 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
pos_Rl (cons t1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons t1 (cons s2 s3)))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
Rlength (cons t1 (cons s2 s3)) = S (Rlength (cons r1 lf'))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons s2 s3)) i) (pos_Rl (cons t1 (cons s2 s3)) (S i))) (pos_Rl (cons r1 lf') i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 lf')))%nat -> pos_Rl (cons r1 lf') i <> pos_Rl (cons r1 lf') (S i) \/ f (pos_Rl (cons t1 (cons s2 s3)) (S i)) <> pos_Rl (cons r1 lf') i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
i:nat
H1:(S i < S (Rlength s3))%nat
Hreci:(i < S (Rlength s3))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <= pos_Rl (cons t1 (cons s2 s3)) (S i)

pos_Rl (cons t1 (cons s2 s3)) (S i) <= pos_Rl (cons t1 (cons s2 s3)) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
pos_Rl (cons t1 (cons s2 s3)) 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
pos_Rl (cons t1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons t1 (cons s2 s3)))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
Rlength (cons t1 (cons s2 s3)) = S (Rlength (cons r1 lf'))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons s2 s3)) i) (pos_Rl (cons t1 (cons s2 s3)) (S i))) (pos_Rl (cons r1 lf') i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 lf')))%nat -> pos_Rl (cons r1 lf') i <> pos_Rl (cons r1 lf') (S i) \/ f (pos_Rl (cons t1 (cons s2 s3)) (S i)) <> pos_Rl (cons r1 lf') i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)

pos_Rl (cons t1 (cons s2 s3)) 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
pos_Rl (cons t1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons t1 (cons s2 s3)))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
Rlength (cons t1 (cons s2 s3)) = S (Rlength (cons r1 lf'))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons s2 s3)) i) (pos_Rl (cons t1 (cons s2 s3)) (S i))) (pos_Rl (cons r1 lf') i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 lf')))%nat -> pos_Rl (cons r1 lf') i <> pos_Rl (cons r1 lf') (S i) \/ f (pos_Rl (cons t1 (cons s2 s3)) (S i)) <> pos_Rl (cons r1 lf') i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)

pos_Rl (cons t1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons t1 (cons s2 s3)))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
Rlength (cons t1 (cons s2 s3)) = S (Rlength (cons r1 lf'))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons s2 s3)) i) (pos_Rl (cons t1 (cons s2 s3)) (S i))) (pos_Rl (cons r1 lf') i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 lf')))%nat -> pos_Rl (cons r1 lf') i <> pos_Rl (cons r1 lf') (S i) \/ f (pos_Rl (cons t1 (cons s2 s3)) (S i)) <> pos_Rl (cons r1 lf') i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)

Rlength (cons t1 (cons s2 s3)) = S (Rlength (cons r1 lf'))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons s2 s3)) i) (pos_Rl (cons t1 (cons s2 s3)) (S i))) (pos_Rl (cons r1 lf') i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 lf')))%nat -> pos_Rl (cons r1 lf') i <> pos_Rl (cons r1 lf') (S i) \/ f (pos_Rl (cons t1 (cons s2 s3)) (S i)) <> pos_Rl (cons r1 lf') i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)

forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons s2 s3)) i) (pos_Rl (cons t1 (cons s2 s3)) (S i))) (pos_Rl (cons r1 lf') i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 lf')))%nat -> pos_Rl (cons r1 lf') i <> pos_Rl (cons r1 lf') (S i) \/ f (pos_Rl (cons t1 (cons s2 s3)) (S i)) <> pos_Rl (cons r1 lf') i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < S (Rlength s3))%nat
x:R
H6:pos_Rl (cons t1 (cons s2 s3)) 0 < x < pos_Rl (cons t1 (cons s2 s3)) 1

f x = pos_Rl (cons r1 lf') 0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
i:nat
H1:(S i < S (Rlength s3))%nat
x:R
H6:pos_Rl (cons t1 (cons s2 s3)) (S i) < x < pos_Rl (cons t1 (cons s2 s3)) (S (S i))
Hreci:(i < S (Rlength s3))%nat -> pos_Rl (cons t1 (cons s2 s3)) i < x < pos_Rl (cons t1 (cons s2 s3)) (S i) -> f x = pos_Rl (cons r1 lf') i
f x = pos_Rl (cons r1 lf') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 lf')))%nat -> pos_Rl (cons r1 lf') i <> pos_Rl (cons r1 lf') (S i) \/ f (pos_Rl (cons t1 (cons s2 s3)) (S i)) <> pos_Rl (cons r1 lf') i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < S (Rlength s3))%nat
x:R
H6:t1 < x < s2
Hlt:x < t2

f x = r1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < S (Rlength s3))%nat
x:R
H6:t1 < x < s2
Heq:x = t2
f x = r1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < S (Rlength s3))%nat
x:R
H6:t1 < x < s2
Hgt:x > t2
f x = r1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
i:nat
H1:(S i < S (Rlength s3))%nat
x:R
H6:pos_Rl (cons t1 (cons s2 s3)) (S i) < x < pos_Rl (cons t1 (cons s2 s3)) (S (S i))
Hreci:(i < S (Rlength s3))%nat -> pos_Rl (cons t1 (cons s2 s3)) i < x < pos_Rl (cons t1 (cons s2 s3)) (S i) -> f x = pos_Rl (cons r1 lf') i
f x = pos_Rl (cons r1 lf') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 lf')))%nat -> pos_Rl (cons r1 lf') i <> pos_Rl (cons r1 lf') (S i) \/ f (pos_Rl (cons t1 (cons s2 s3)) (S i)) <> pos_Rl (cons r1 lf') i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < S (Rlength s3))%nat
x:R
H6:t1 < x < s2
Heq:x = t2

f x = r1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < S (Rlength s3))%nat
x:R
H6:t1 < x < s2
Hgt:x > t2
f x = r1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
i:nat
H1:(S i < S (Rlength s3))%nat
x:R
H6:pos_Rl (cons t1 (cons s2 s3)) (S i) < x < pos_Rl (cons t1 (cons s2 s3)) (S (S i))
Hreci:(i < S (Rlength s3))%nat -> pos_Rl (cons t1 (cons s2 s3)) i < x < pos_Rl (cons t1 (cons s2 s3)) (S i) -> f x = pos_Rl (cons r1 lf') i
f x = pos_Rl (cons r1 lf') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 lf')))%nat -> pos_Rl (cons r1 lf') i <> pos_Rl (cons r1 lf') (S i) \/ f (pos_Rl (cons t1 (cons s2 s3)) (S i)) <> pos_Rl (cons r1 lf') i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < S (Rlength s3))%nat
x:R
H6:t1 < x < s2
Hgt:x > t2

f x = r1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
i:nat
H1:(S i < S (Rlength s3))%nat
x:R
H6:pos_Rl (cons t1 (cons s2 s3)) (S i) < x < pos_Rl (cons t1 (cons s2 s3)) (S (S i))
Hreci:(i < S (Rlength s3))%nat -> pos_Rl (cons t1 (cons s2 s3)) i < x < pos_Rl (cons t1 (cons s2 s3)) (S i) -> f x = pos_Rl (cons r1 lf') i
f x = pos_Rl (cons r1 lf') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 lf')))%nat -> pos_Rl (cons r1 lf') i <> pos_Rl (cons r1 lf') (S i) \/ f (pos_Rl (cons t1 (cons s2 s3)) (S i)) <> pos_Rl (cons r1 lf') i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
i:nat
H1:(S i < S (Rlength s3))%nat
x:R
H6:pos_Rl (cons t1 (cons s2 s3)) (S i) < x < pos_Rl (cons t1 (cons s2 s3)) (S (S i))
Hreci:(i < S (Rlength s3))%nat -> pos_Rl (cons t1 (cons s2 s3)) i < x < pos_Rl (cons t1 (cons s2 s3)) (S i) -> f x = pos_Rl (cons r1 lf') i

f x = pos_Rl (cons r1 lf') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 lf')))%nat -> pos_Rl (cons r1 lf') i <> pos_Rl (cons r1 lf') (S i) \/ f (pos_Rl (cons t1 (cons s2 s3)) (S i)) <> pos_Rl (cons r1 lf') i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)

forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 lf')))%nat -> pos_Rl (cons r1 lf') i <> pos_Rl (cons r1 lf') (S i) \/ f (pos_Rl (cons t1 (cons s2 s3)) (S i)) <> pos_Rl (cons r1 lf') i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)

forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat

pos_Rl (cons t1 (cons s2 s3)) 0 <> pos_Rl (cons t1 (cons s2 s3)) 1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat
Hreci:(i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
pos_Rl (cons t1 (cons s2 s3)) (S i) <> pos_Rl (cons t1 (cons s2 s3)) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat
H6:t1 = s2

t1 <= t2
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat
H6:t1 = s2
t2 <= t1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat
Hreci:(i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
pos_Rl (cons t1 (cons s2 s3)) (S i) <> pos_Rl (cons t1 (cons s2 s3)) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H1:(0 < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat
H6:t1 = s2

t2 <= t1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat
Hreci:(i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)
pos_Rl (cons t1 (cons s2 s3)) (S i) <> pos_Rl (cons t1 (cons s2 s3)) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 = r1
H12:ordered_Rlist (cons t1 (cons t2 t3))
H14:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H13:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H15:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
H16:ordered_Rlist (cons s1 (cons s2 s3))
H19:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H18:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H20:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat
Hreci:(i < Init.Nat.pred (Rlength (cons t1 (cons s2 s3))))%nat -> pos_Rl (cons t1 (cons s2 s3)) i <> pos_Rl (cons t1 (cons s2 s3)) (S i)

pos_Rl (cons t1 (cons s2 s3)) (S i) <> pos_Rl (cons t1 (cons s2 s3)) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1

exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)

ordered_Rlist (cons t1 l')
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') (Init.Nat.pred (Rlength (cons t1 l'))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
Rlength (cons t1 l') = S (Rlength (cons r1 (cons r2 lf')))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 l') i) (pos_Rl (cons t1 l') (S i))) (pos_Rl (cons r1 (cons r2 lf')) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < S (S (Rlength s3)))%nat

pos_Rl (cons t1 (cons s1 (cons s2 s3))) 0 <= pos_Rl (cons t1 (cons s1 (cons s2 s3))) 1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < S (S (Rlength s3)))%nat
Hreci:(i < S (S (Rlength s3)))%nat -> pos_Rl (cons t1 (cons s1 (cons s2 s3))) i <= pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i)
pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i) <= pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') (Init.Nat.pred (Rlength (cons t1 l'))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
Rlength (cons t1 l') = S (Rlength (cons r1 (cons r2 lf')))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 l') i) (pos_Rl (cons t1 l') (S i))) (pos_Rl (cons r1 (cons r2 lf')) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < S (S (Rlength s3)))%nat

t1 <= t2
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < S (S (Rlength s3)))%nat
t2 = s1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < S (S (Rlength s3)))%nat
Hreci:(i < S (S (Rlength s3)))%nat -> pos_Rl (cons t1 (cons s1 (cons s2 s3))) i <= pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i)
pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i) <= pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') (Init.Nat.pred (Rlength (cons t1 l'))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
Rlength (cons t1 l') = S (Rlength (cons r1 (cons r2 lf')))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 l') i) (pos_Rl (cons t1 l') (S i))) (pos_Rl (cons r1 (cons r2 lf')) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < S (S (Rlength s3)))%nat

t2 = s1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < S (S (Rlength s3)))%nat
Hreci:(i < S (S (Rlength s3)))%nat -> pos_Rl (cons t1 (cons s1 (cons s2 s3))) i <= pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i)
pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i) <= pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') (Init.Nat.pred (Rlength (cons t1 l'))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
Rlength (cons t1 l') = S (Rlength (cons r1 (cons r2 lf')))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 l') i) (pos_Rl (cons t1 l') (S i))) (pos_Rl (cons r1 (cons r2 lf')) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < S (S (Rlength s3)))%nat
Hreci:(i < S (S (Rlength s3)))%nat -> pos_Rl (cons t1 (cons s1 (cons s2 s3))) i <= pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i)

pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i) <= pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') (Init.Nat.pred (Rlength (cons t1 l'))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
Rlength (cons t1 l') = S (Rlength (cons r1 (cons r2 lf')))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 l') i) (pos_Rl (cons t1 l') (S i))) (pos_Rl (cons r1 (cons r2 lf')) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)

pos_Rl (cons t1 l') 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') (Init.Nat.pred (Rlength (cons t1 l'))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
Rlength (cons t1 l') = S (Rlength (cons r1 (cons r2 lf')))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 l') i) (pos_Rl (cons t1 l') (S i))) (pos_Rl (cons r1 (cons r2 lf')) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)

pos_Rl (cons t1 l') (Init.Nat.pred (Rlength (cons t1 l'))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
Rlength (cons t1 l') = S (Rlength (cons r1 (cons r2 lf')))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 l') i) (pos_Rl (cons t1 l') (S i))) (pos_Rl (cons r1 (cons r2 lf')) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)

Rlength (cons t1 l') = S (Rlength (cons r1 (cons r2 lf')))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 l') i) (pos_Rl (cons t1 l') (S i))) (pos_Rl (cons r1 (cons r2 lf')) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)

forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 l') i) (pos_Rl (cons t1 l') (S i))) (pos_Rl (cons r1 (cons r2 lf')) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Rlength l')%nat
x:R
H6:pos_Rl (cons t1 l') 0 < x < pos_Rl (cons t1 l') 1

f x = pos_Rl (cons r1 (cons r2 lf')) 0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Rlength l')%nat
x:R
H6:pos_Rl (cons t1 l') (S i) < x < pos_Rl (cons t1 l') (S (S i))
Hreci:(i < Rlength l')%nat -> pos_Rl (cons t1 l') i < x < pos_Rl (cons t1 l') (S i) -> f x = pos_Rl (cons r1 (cons r2 lf')) i
f x = pos_Rl (cons r1 (cons r2 lf')) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Rlength l')%nat
x:R
H6:t1 < x < s1

(0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Rlength l')%nat
x:R
H6:t1 < x < s1
open_interval (pos_Rl (cons t1 (cons t2 t3)) 0) (pos_Rl (cons t1 (cons t2 t3)) 1) x
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Rlength l')%nat
x:R
H6:pos_Rl (cons t1 l') (S i) < x < pos_Rl (cons t1 l') (S (S i))
Hreci:(i < Rlength l')%nat -> pos_Rl (cons t1 l') i < x < pos_Rl (cons t1 l') (S i) -> f x = pos_Rl (cons r1 (cons r2 lf')) i
f x = pos_Rl (cons r1 (cons r2 lf')) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Rlength l')%nat
x:R
H6:t1 < x < s1

open_interval (pos_Rl (cons t1 (cons t2 t3)) 0) (pos_Rl (cons t1 (cons t2 t3)) 1) x
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Rlength l')%nat
x:R
H6:pos_Rl (cons t1 l') (S i) < x < pos_Rl (cons t1 l') (S (S i))
Hreci:(i < Rlength l')%nat -> pos_Rl (cons t1 l') i < x < pos_Rl (cons t1 l') (S i) -> f x = pos_Rl (cons r1 (cons r2 lf')) i
f x = pos_Rl (cons r1 (cons r2 lf')) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Rlength l')%nat
x:R
H6:t1 < x < s1

t1 < x < t2
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Rlength l')%nat
x:R
H6:pos_Rl (cons t1 l') (S i) < x < pos_Rl (cons t1 l') (S (S i))
Hreci:(i < Rlength l')%nat -> pos_Rl (cons t1 l') i < x < pos_Rl (cons t1 l') (S i) -> f x = pos_Rl (cons r1 (cons r2 lf')) i
f x = pos_Rl (cons r1 (cons r2 lf')) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Rlength l')%nat
x:R
H6:t1 < x < s1

t1 < x < s1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Rlength l')%nat
x:R
H6:t1 < x < s1
s1 = t2
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Rlength l')%nat
x:R
H6:pos_Rl (cons t1 l') (S i) < x < pos_Rl (cons t1 l') (S (S i))
Hreci:(i < Rlength l')%nat -> pos_Rl (cons t1 l') i < x < pos_Rl (cons t1 l') (S i) -> f x = pos_Rl (cons r1 (cons r2 lf')) i
f x = pos_Rl (cons r1 (cons r2 lf')) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Rlength l')%nat
x:R
H6:t1 < x < s1

s1 = t2
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Rlength l')%nat
x:R
H6:pos_Rl (cons t1 l') (S i) < x < pos_Rl (cons t1 l') (S (S i))
Hreci:(i < Rlength l')%nat -> pos_Rl (cons t1 l') i < x < pos_Rl (cons t1 l') (S i) -> f x = pos_Rl (cons r1 (cons r2 lf')) i
f x = pos_Rl (cons r1 (cons r2 lf')) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Rlength l')%nat
x:R
H6:pos_Rl (cons t1 l') (S i) < x < pos_Rl (cons t1 l') (S (S i))
Hreci:(i < Rlength l')%nat -> pos_Rl (cons t1 l') i < x < pos_Rl (cons t1 l') (S i) -> f x = pos_Rl (cons r1 (cons r2 lf')) i

f x = pos_Rl (cons r1 (cons r2 lf')) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Rlength l')%nat
x:R
H6:pos_Rl (cons t1 l') (S i) < x < pos_Rl (cons t1 l') (S (S i))

(i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Rlength l')%nat
x:R
H6:pos_Rl (cons t1 l') (S i) < x < pos_Rl (cons t1 l') (S (S i))
open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i)) x
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Rlength l')%nat
x:R
H6:pos_Rl (cons t1 l') (S i) < x < pos_Rl (cons t1 l') (S (S i))

open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i)) x
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)

forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < S (Rlength lf'))%nat

pos_Rl (cons r1 (cons r2 lf')) 0 <> pos_Rl (cons r1 (cons r2 lf')) 1 \/ f (pos_Rl (cons t1 l') 1) <> pos_Rl (cons r1 (cons r2 lf')) 0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < S (Rlength lf'))%nat
Hreci:(i < S (Rlength lf'))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
pos_Rl (cons r1 (cons r2 lf')) (S i) <> pos_Rl (cons r1 (cons r2 lf')) (S (S i)) \/ f (pos_Rl (cons t1 l') (S (S i))) <> pos_Rl (cons r1 (cons r2 lf')) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < S (Rlength lf'))%nat

f t2 <> r1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < S (Rlength lf'))%nat
t2 = s1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < S (Rlength lf'))%nat
Hreci:(i < S (Rlength lf'))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
pos_Rl (cons r1 (cons r2 lf')) (S i) <> pos_Rl (cons r1 (cons r2 lf')) (S (S i)) \/ f (pos_Rl (cons t1 l') (S (S i))) <> pos_Rl (cons r1 (cons r2 lf')) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < S (Rlength lf'))%nat

t2 = s1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < S (Rlength lf'))%nat
Hreci:(i < S (Rlength lf'))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
pos_Rl (cons r1 (cons r2 lf')) (S i) <> pos_Rl (cons r1 (cons r2 lf')) (S (S i)) \/ f (pos_Rl (cons t1 l') (S (S i))) <> pos_Rl (cons r1 (cons r2 lf')) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < S (Rlength lf'))%nat
Hreci:(i < S (Rlength lf'))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i

pos_Rl (cons r1 (cons r2 lf')) (S i) <> pos_Rl (cons r1 (cons r2 lf')) (S (S i)) \/ f (pos_Rl (cons t1 l') (S (S i))) <> pos_Rl (cons r1 (cons r2 lf')) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < S (Rlength lf'))%nat
Hreci:(i < S (Rlength lf'))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0)

(i < Init.Nat.pred (Rlength (cons r2 lf')))%nat
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)

forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Init.Nat.pred (Rlength (cons t1 l')))%nat

pos_Rl (cons t1 (cons s1 (cons s2 s3))) 0 <> pos_Rl (cons t1 (cons s1 (cons s2 s3))) 1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons t1 l')))%nat
Hreci:(i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 (cons s1 (cons s2 s3))) i <> pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i)
pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i) <> pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Init.Nat.pred (Rlength (cons t1 l')))%nat
H6:t1 = s1

t1 <= t2
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Init.Nat.pred (Rlength (cons t1 l')))%nat
H6:t1 = s1
t2 <= t1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons t1 l')))%nat
Hreci:(i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 (cons s1 (cons s2 s3))) i <> pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i)
pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i) <> pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Init.Nat.pred (Rlength (cons t1 l')))%nat
H6:t1 = s1

t2 <= t1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons t1 l')))%nat
Hreci:(i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 (cons s1 (cons s2 s3))) i <> pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i)
pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i) <> pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 = r2
H11:f t2 <> r1
H12:ordered_Rlist (cons s1 (cons s2 s3))
H14:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H13:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H15:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H17:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H16:ordered_Rlist (cons t1 (cons t2 t3))
H19:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H18:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H20:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H22:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons t1 l')))%nat
Hreci:(i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 (cons s1 (cons s2 s3))) i <> pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i)

pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i) <> pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
H6:adapted_couple f t2 b l' (cons r2 lf')
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
s1, s2:R
s3:Rlist
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2

exists l'0 lf'0 : Rlist, adapted_couple_opt f a b l'0 lf'0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)

ordered_Rlist (cons t1 l')
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') (Init.Nat.pred (Rlength (cons t1 l'))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
Rlength (cons t1 l') = S (Rlength (cons r1 (cons r2 lf')))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 l') i) (pos_Rl (cons t1 l') (S i))) (pos_Rl (cons r1 (cons r2 lf')) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < S (S (Rlength s3)))%nat

pos_Rl (cons t1 (cons s1 (cons s2 s3))) 0 <= pos_Rl (cons t1 (cons s1 (cons s2 s3))) 1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < S (S (Rlength s3)))%nat
Hreci:(i < S (S (Rlength s3)))%nat -> pos_Rl (cons t1 (cons s1 (cons s2 s3))) i <= pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i)
pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i) <= pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') (Init.Nat.pred (Rlength (cons t1 l'))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
Rlength (cons t1 l') = S (Rlength (cons r1 (cons r2 lf')))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 l') i) (pos_Rl (cons t1 l') (S i))) (pos_Rl (cons r1 (cons r2 lf')) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < S (S (Rlength s3)))%nat

t1 <= t2
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < S (S (Rlength s3)))%nat
t2 = s1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < S (S (Rlength s3)))%nat
Hreci:(i < S (S (Rlength s3)))%nat -> pos_Rl (cons t1 (cons s1 (cons s2 s3))) i <= pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i)
pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i) <= pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') (Init.Nat.pred (Rlength (cons t1 l'))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
Rlength (cons t1 l') = S (Rlength (cons r1 (cons r2 lf')))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 l') i) (pos_Rl (cons t1 l') (S i))) (pos_Rl (cons r1 (cons r2 lf')) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < S (S (Rlength s3)))%nat

t2 = s1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < S (S (Rlength s3)))%nat
Hreci:(i < S (S (Rlength s3)))%nat -> pos_Rl (cons t1 (cons s1 (cons s2 s3))) i <= pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i)
pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i) <= pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') (Init.Nat.pred (Rlength (cons t1 l'))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
Rlength (cons t1 l') = S (Rlength (cons r1 (cons r2 lf')))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 l') i) (pos_Rl (cons t1 l') (S i))) (pos_Rl (cons r1 (cons r2 lf')) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < S (S (Rlength s3)))%nat
Hreci:(i < S (S (Rlength s3)))%nat -> pos_Rl (cons t1 (cons s1 (cons s2 s3))) i <= pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i)

pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i) <= pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') (Init.Nat.pred (Rlength (cons t1 l'))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
Rlength (cons t1 l') = S (Rlength (cons r1 (cons r2 lf')))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 l') i) (pos_Rl (cons t1 l') (S i))) (pos_Rl (cons r1 (cons r2 lf')) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)

pos_Rl (cons t1 l') 0 = Rmin a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
pos_Rl (cons t1 l') (Init.Nat.pred (Rlength (cons t1 l'))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
Rlength (cons t1 l') = S (Rlength (cons r1 (cons r2 lf')))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 l') i) (pos_Rl (cons t1 l') (S i))) (pos_Rl (cons r1 (cons r2 lf')) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)

pos_Rl (cons t1 l') (Init.Nat.pred (Rlength (cons t1 l'))) = Rmax a b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
Rlength (cons t1 l') = S (Rlength (cons r1 (cons r2 lf')))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 l') i) (pos_Rl (cons t1 l') (S i))) (pos_Rl (cons r1 (cons r2 lf')) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)

Rlength (cons t1 l') = S (Rlength (cons r1 (cons r2 lf')))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 l') i) (pos_Rl (cons t1 l') (S i))) (pos_Rl (cons r1 (cons r2 lf')) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)

forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 l') i) (pos_Rl (cons t1 l') (S i))) (pos_Rl (cons r1 (cons r2 lf')) i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Rlength l')%nat
x:R
H6:pos_Rl (cons t1 l') 0 < x < pos_Rl (cons t1 l') 1

f x = pos_Rl (cons r1 (cons r2 lf')) 0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Rlength l')%nat
x:R
H6:pos_Rl (cons t1 l') (S i) < x < pos_Rl (cons t1 l') (S (S i))
Hreci:(i < Rlength l')%nat -> pos_Rl (cons t1 l') i < x < pos_Rl (cons t1 l') (S i) -> f x = pos_Rl (cons r1 (cons r2 lf')) i
f x = pos_Rl (cons r1 (cons r2 lf')) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Rlength l')%nat
x:R
H6:t1 < x < s1

(0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Rlength l')%nat
x:R
H6:t1 < x < s1
open_interval (pos_Rl (cons t1 (cons t2 t3)) 0) (pos_Rl (cons t1 (cons t2 t3)) 1) x
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Rlength l')%nat
x:R
H6:pos_Rl (cons t1 l') (S i) < x < pos_Rl (cons t1 l') (S (S i))
Hreci:(i < Rlength l')%nat -> pos_Rl (cons t1 l') i < x < pos_Rl (cons t1 l') (S i) -> f x = pos_Rl (cons r1 (cons r2 lf')) i
f x = pos_Rl (cons r1 (cons r2 lf')) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Rlength l')%nat
x:R
H6:t1 < x < s1

open_interval (pos_Rl (cons t1 (cons t2 t3)) 0) (pos_Rl (cons t1 (cons t2 t3)) 1) x
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Rlength l')%nat
x:R
H6:pos_Rl (cons t1 l') (S i) < x < pos_Rl (cons t1 l') (S (S i))
Hreci:(i < Rlength l')%nat -> pos_Rl (cons t1 l') i < x < pos_Rl (cons t1 l') (S i) -> f x = pos_Rl (cons r1 (cons r2 lf')) i
f x = pos_Rl (cons r1 (cons r2 lf')) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Rlength l')%nat
x:R
H6:t1 < x < s1

t1 < x < s1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Rlength l')%nat
x:R
H6:t1 < x < s1
s1 = t2
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Rlength l')%nat
x:R
H6:pos_Rl (cons t1 l') (S i) < x < pos_Rl (cons t1 l') (S (S i))
Hreci:(i < Rlength l')%nat -> pos_Rl (cons t1 l') i < x < pos_Rl (cons t1 l') (S i) -> f x = pos_Rl (cons r1 (cons r2 lf')) i
f x = pos_Rl (cons r1 (cons r2 lf')) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Rlength l')%nat
x:R
H6:t1 < x < s1

s1 = t2
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Rlength l')%nat
x:R
H6:pos_Rl (cons t1 l') (S i) < x < pos_Rl (cons t1 l') (S (S i))
Hreci:(i < Rlength l')%nat -> pos_Rl (cons t1 l') i < x < pos_Rl (cons t1 l') (S i) -> f x = pos_Rl (cons r1 (cons r2 lf')) i
f x = pos_Rl (cons r1 (cons r2 lf')) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Rlength l')%nat
x:R
H6:pos_Rl (cons t1 l') (S i) < x < pos_Rl (cons t1 l') (S (S i))
Hreci:(i < Rlength l')%nat -> pos_Rl (cons t1 l') i < x < pos_Rl (cons t1 l') (S i) -> f x = pos_Rl (cons r1 (cons r2 lf')) i

f x = pos_Rl (cons r1 (cons r2 lf')) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Rlength l')%nat
x:R
H6:pos_Rl (cons t1 l') (S i) < x < pos_Rl (cons t1 l') (S (S i))

(i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Rlength l')%nat
x:R
H6:pos_Rl (cons t1 l') (S i) < x < pos_Rl (cons t1 l') (S (S i))
open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i)) x
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Rlength l')%nat
x:R
H6:pos_Rl (cons t1 l') (S i) < x < pos_Rl (cons t1 l') (S (S i))

open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i)) x
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)

forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons r2 lf'))))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < S (Rlength lf'))%nat

pos_Rl (cons r1 (cons r2 lf')) 0 <> pos_Rl (cons r1 (cons r2 lf')) 1 \/ f (pos_Rl (cons t1 l') 1) <> pos_Rl (cons r1 (cons r2 lf')) 0
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < S (Rlength lf'))%nat
Hreci:(i < S (Rlength lf'))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
pos_Rl (cons r1 (cons r2 lf')) (S i) <> pos_Rl (cons r1 (cons r2 lf')) (S (S i)) \/ f (pos_Rl (cons t1 l') (S (S i))) <> pos_Rl (cons r1 (cons r2 lf')) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < S (Rlength lf'))%nat
Hreci:(i < S (Rlength lf'))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i

pos_Rl (cons r1 (cons r2 lf')) (S i) <> pos_Rl (cons r1 (cons r2 lf')) (S (S i)) \/ f (pos_Rl (cons t1 l') (S (S i))) <> pos_Rl (cons r1 (cons r2 lf')) (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < S (Rlength lf'))%nat
Hreci:(i < S (Rlength lf'))%nat -> pos_Rl (cons r1 (cons r2 lf')) i <> pos_Rl (cons r1 (cons r2 lf')) (S i) \/ f (pos_Rl (cons t1 l') (S i)) <> pos_Rl (cons r1 (cons r2 lf')) i
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0
H20:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0)

(i < Init.Nat.pred (Rlength (cons r2 lf')))%nat
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)

forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 l') i <> pos_Rl (cons t1 l') (S i)
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Init.Nat.pred (Rlength (cons t1 l')))%nat

pos_Rl (cons t1 (cons s1 (cons s2 s3))) 0 <> pos_Rl (cons t1 (cons s1 (cons s2 s3))) 1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons t1 l')))%nat
Hreci:(i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 (cons s1 (cons s2 s3))) i <> pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i)
pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i) <> pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Init.Nat.pred (Rlength (cons t1 l')))%nat
H6:t1 = s1

t1 <= t2
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Init.Nat.pred (Rlength (cons t1 l')))%nat
H6:t1 = s1
t2 <= t1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons t1 l')))%nat
Hreci:(i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 (cons s1 (cons s2 s3))) i <> pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i)
pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i) <> pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i <> pos_Rl (cons r2 lf') (S i) \/ f (pos_Rl l' (S i)) <> pos_Rl (cons r2 lf') i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i <> pos_Rl l' (S i))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r2 lf') i)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i) (pos_Rl (cons t1 (cons t2 t3)) (S i))) (pos_Rl (cons r1 lf) i)
H1:(0 < Init.Nat.pred (Rlength (cons t1 l')))%nat
H6:t1 = s1

t2 <= t1
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons t1 l')))%nat
Hreci:(i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 (cons s1 (cons s2 s3))) i <> pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i)
pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i) <> pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l'0 lf'0 : Rlist, adapted_couple_opt f a0 b0 l'0 lf'0
r1:R
lf:Rlist
a, b:R
H0:a <= b
t1, t2:R
t3:Rlist
H2:a <> b
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
H5:t2 <= b
Hyp_eq:t1 <> t2
l':Rlist
r2:R
lf':Rlist
H7:t2 <> b
Hyp_min:Rmin t2 b = t2
s1, s2:R
s3:Rlist
H8:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r2 lf')))%nat -> pos_Rl (cons r2 lf') i0 <> pos_Rl (cons r2 lf') (S i0) \/ f (pos_Rl l' (S i0)) <> pos_Rl (cons r2 lf') i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l'))%nat -> pos_Rl l' i0 <> pos_Rl l' (S i0))
H9:l' = cons s1 (cons s2 s3)
H10:r1 <> r2
H11:ordered_Rlist (cons s1 (cons s2 s3))
H13:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin t2 b
H12:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax t2 b
H14:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r2 lf'))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r2 lf') i0)
H15:ordered_Rlist (cons t1 (cons t2 t3))
H18:pos_Rl (cons t1 (cons t2 t3)) 0 = Rmin a b
H17:pos_Rl (cons t1 (cons t2 t3)) (Init.Nat.pred (Rlength (cons t1 (cons t2 t3)))) = Rmax a b
H19:Rlength (cons t1 (cons t2 t3)) = S (Rlength (cons r1 lf))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons t1 (cons t2 t3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons t1 (cons t2 t3)) i0) (pos_Rl (cons t1 (cons t2 t3)) (S i0))) (pos_Rl (cons r1 lf) i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons t1 l')))%nat
Hreci:(i < Init.Nat.pred (Rlength (cons t1 l')))%nat -> pos_Rl (cons t1 (cons s1 (cons s2 s3))) i <> pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i)

pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S i) <> pos_Rl (cons t1 (cons s1 (cons s2 s3))) (S (S i))
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf
t2 <= b
f:R -> R
l:Rlist
r:R
r0:Rlist
H:forall (lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 r0 lf0 -> exists l' lf' : Rlist, adapted_couple_opt f a0 b0 l' lf'
r1:R
lf:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r r0) (cons r1 lf)
H2:a <> b
t1, t2:R
t3:Rlist
H3:cons r r0 = cons t1 (cons t2 t3)
H4:adapted_couple f t2 b r0 lf

t2 <= b
rewrite H3 in H1; clear H4; unfold adapted_couple in H1; decompose [and] H1; clear H1; clear H H7 H9; cut (Rmax a b = b); [ intro; rewrite H in H5; rewrite <- H5; apply RList_P7; [ assumption | simpl; right; left; reflexivity ] | unfold Rmax; decide (Rle_dec a b) with H0; reflexivity ]. Qed.

forall (a b r r1 r3 s1 s2 r4 : R) (r2 lf1 s3 lf2 : Rlist) (f : R -> R), a < b -> adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1) -> adapted_couple_opt f a b (cons s1 (cons s2 s3)) (cons r4 lf2) -> r1 <= s2

forall (a b r r1 r3 s1 s2 r4 : R) (r2 lf1 s3 lf2 : Rlist) (f : R -> R), a < b -> adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1) -> adapted_couple_opt f a b (cons s1 (cons s2 s3)) (cons r4 lf2) -> r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)

r = s1
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1

r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1

r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 < s2

r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2

r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1

r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 nil)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 nil))))%nat -> pos_Rl (cons s1 (cons s2 nil)) i <> pos_Rl (cons s1 (cons s2 nil)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 nil))
H10:pos_Rl (cons s1 (cons s2 nil)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 nil)) (Init.Nat.pred (Rlength (cons s1 (cons s2 nil)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 nil)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 nil))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 nil)) i) (pos_Rl (cons s1 (cons s2 nil)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1

r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
Hrecs3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i)) -> ordered_Rlist (cons s1 (cons s2 s3)) -> pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b -> pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b -> Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2)) -> (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)) -> r1 <= s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 nil)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 nil))))%nat -> pos_Rl (cons s1 (cons s2 nil)) i <> pos_Rl (cons s1 (cons s2 nil)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 nil))
H10:pos_Rl (cons s1 (cons s2 nil)) 0 = Rmin a b
H9:s2 = Rmax a b
H11:Rlength (cons s1 (cons s2 nil)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 nil))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 nil)) i) (pos_Rl (cons s1 (cons s2 nil)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:Rmax a b < r1

r1 <= Rmax a b -> r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 nil)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 nil))))%nat -> pos_Rl (cons s1 (cons s2 nil)) i <> pos_Rl (cons s1 (cons s2 nil)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 nil))
H10:pos_Rl (cons s1 (cons s2 nil)) 0 = Rmin a b
H9:s2 = Rmax a b
H11:Rlength (cons s1 (cons s2 nil)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 nil))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 nil)) i) (pos_Rl (cons s1 (cons s2 nil)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:Rmax a b < r1
r1 <= Rmax a b
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
Hrecs3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i)) -> ordered_Rlist (cons s1 (cons s2 s3)) -> pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b -> pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b -> Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2)) -> (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)) -> r1 <= s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 nil)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 nil))))%nat -> pos_Rl (cons s1 (cons s2 nil)) i <> pos_Rl (cons s1 (cons s2 nil)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 nil))
H10:pos_Rl (cons s1 (cons s2 nil)) 0 = Rmin a b
H9:s2 = Rmax a b
H11:Rlength (cons s1 (cons s2 nil)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 nil))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 nil)) i) (pos_Rl (cons s1 (cons s2 nil)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:Rmax a b < r1

r1 <= Rmax a b
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
Hrecs3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i)) -> ordered_Rlist (cons s1 (cons s2 s3)) -> pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b -> pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b -> Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2)) -> (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)) -> r1 <= s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
Hrecs3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i)) -> ordered_Rlist (cons s1 (cons s2 s3)) -> pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b -> pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b -> Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2)) -> (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)) -> r1 <= s2

r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 nil)))%nat -> pos_Rl (cons r4 nil) i <> pos_Rl (cons r4 nil) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 nil) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 nil))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 nil) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1

r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
Hreclf2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) -> Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 lf2)) -> (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 lf2) i)) -> r1 <= s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
Hreclf2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) -> Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 lf2)) -> (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 lf2) i)) -> r1 <= s2

r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1

r3 = r4
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
x:=(r + s2) / 2:R
H17:forall x0 : R, r < x0 < r1 -> f x0 = r3
H18:forall x0 : R, s1 < x0 < s2 -> f x0 = r4

f x = r4
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
x:=(r + s2) / 2:R
H17:forall x0 : R, r < x0 < r1 -> f x0 = r3
H18:forall x0 : R, s1 < x0 < s2 -> f x0 = r4
r < x < r1
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
x:=(r + s2) / 2:R
H17:forall x0 : R, r < x0 < r1 -> f x0 = r3
H18:forall x0 : R, s1 < x0 < s2 -> f x0 = r4

f x = f x
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
x:=(r + s2) / 2:R
H17:forall x0 : R, r < x0 < r1 -> f x0 = r3
H18:forall x0 : R, s1 < x0 < s2 -> f x0 = r4
s1 < x < s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
x:=(r + s2) / 2:R
H17:forall x0 : R, r < x0 < r1 -> f x0 = r3
H18:forall x0 : R, s1 < x0 < s2 -> f x0 = r4
r < x < r1
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
x:=(r + s2) / 2:R
H17:forall x0 : R, r < x0 < r1 -> f x0 = r3
H18:forall x0 : R, s1 < x0 < s2 -> f x0 = r4

s1 < x < s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
x:=(r + s2) / 2:R
H17:forall x0 : R, r < x0 < r1 -> f x0 = r3
H18:forall x0 : R, s1 < x0 < s2 -> f x0 = r4
r < x < r1
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
x:=(r + s2) / 2:R
H17:forall x0 : R, r < x0 < r1 -> f x0 = r3
H18:forall x0 : R, s1 < x0 < s2 -> f x0 = r4

r < (r + s2) / 2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
x:=(r + s2) / 2:R
H17:forall x0 : R, r < x0 < r1 -> f x0 = r3
H18:forall x0 : R, s1 < x0 < s2 -> f x0 = r4
(r + s2) / 2 < s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
x:=(r + s2) / 2:R
H17:forall x0 : R, r < x0 < r1 -> f x0 = r3
H18:forall x0 : R, s1 < x0 < s2 -> f x0 = r4
r < x < r1
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
x:=(r + s2) / 2:R
H17:forall x0 : R, r < x0 < r1 -> f x0 = r3
H18:forall x0 : R, s1 < x0 < s2 -> f x0 = r4

(r + s2) / 2 < s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
x:=(r + s2) / 2:R
H17:forall x0 : R, r < x0 < r1 -> f x0 = r3
H18:forall x0 : R, s1 < x0 < s2 -> f x0 = r4
r < x < r1
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
x:=(r + s2) / 2:R
H17:forall x0 : R, r < x0 < r1 -> f x0 = r3
H18:forall x0 : R, s1 < x0 < s2 -> f x0 = r4

r < x < r1
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
x:=(r + s2) / 2:R
H17:forall x0 : R, r < x0 < r1 -> f x0 = r3
H18:forall x0 : R, s1 < x0 < s2 -> f x0 = r4

r < (r + s2) / 2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
x:=(r + s2) / 2:R
H17:forall x0 : R, r < x0 < r1 -> f x0 = r3
H18:forall x0 : R, s1 < x0 < s2 -> f x0 = r4
(r + s2) / 2 < r1
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
x:=(r + s2) / 2:R
H17:forall x0 : R, r < x0 < r1 -> f x0 = r3
H18:forall x0 : R, s1 < x0 < s2 -> f x0 = r4

(r + s2) / 2 < r1
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4

r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4

f s2 = r3
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3

r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3

r3 = r5
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:r3 = r5
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 < r0

r3 = r5
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 = r0
r3 = r5
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:r3 = r5
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 < r0
x:=(s2 + Rmin r1 r0) / 2:R
H22:(0 < S (Rlength r2))%nat -> constant_D_eq f (open_interval r r1) r3
H23:(1 < S (S (Rlength s3)))%nat -> constant_D_eq f (open_interval s2 r0) r5

f x = r5
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 < r0
x:=(s2 + Rmin r1 r0) / 2:R
H22:(0 < S (Rlength r2))%nat -> constant_D_eq f (open_interval r r1) r3
H23:(1 < S (S (Rlength s3)))%nat -> constant_D_eq f (open_interval s2 r0) r5
open_interval r r1 x
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 = r0
r3 = r5
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:r3 = r5
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 < r0
x:=(s2 + Rmin r1 r0) / 2:R
H22:(0 < S (Rlength r2))%nat -> constant_D_eq f (open_interval r r1) r3
H23:(1 < S (S (Rlength s3)))%nat -> constant_D_eq f (open_interval s2 r0) r5

f x = f x
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 < r0
x:=(s2 + Rmin r1 r0) / 2:R
H22:(0 < S (Rlength r2))%nat -> constant_D_eq f (open_interval r r1) r3
H23:(1 < S (S (Rlength s3)))%nat -> constant_D_eq f (open_interval s2 r0) r5
open_interval s2 r0 x
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 < r0
x:=(s2 + Rmin r1 r0) / 2:R
H22:(0 < S (Rlength r2))%nat -> constant_D_eq f (open_interval r r1) r3
H23:(1 < S (S (Rlength s3)))%nat -> constant_D_eq f (open_interval s2 r0) r5
open_interval r r1 x
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 = r0
r3 = r5
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:r3 = r5
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 < r0
x:=(s2 + Rmin r1 r0) / 2:R
H22:(0 < S (Rlength r2))%nat -> constant_D_eq f (open_interval r r1) r3
H23:(1 < S (S (Rlength s3)))%nat -> constant_D_eq f (open_interval s2 r0) r5

open_interval s2 r0 x
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 < r0
x:=(s2 + Rmin r1 r0) / 2:R
H22:(0 < S (Rlength r2))%nat -> constant_D_eq f (open_interval r r1) r3
H23:(1 < S (S (Rlength s3)))%nat -> constant_D_eq f (open_interval s2 r0) r5
open_interval r r1 x
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 = r0
r3 = r5
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:r3 = r5
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 < r0
x:=(s2 + Rmin r1 r0) / 2:R
H22:(0 < S (Rlength r2))%nat -> constant_D_eq f (open_interval r r1) r3
H23:(1 < S (S (Rlength s3)))%nat -> constant_D_eq f (open_interval s2 r0) r5

s2 < (s2 + Rmin r1 r0) / 2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 < r0
x:=(s2 + Rmin r1 r0) / 2:R
H22:(0 < S (Rlength r2))%nat -> constant_D_eq f (open_interval r r1) r3
H23:(1 < S (S (Rlength s3)))%nat -> constant_D_eq f (open_interval s2 r0) r5
(s2 + Rmin r1 r0) / 2 < r0
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 < r0
x:=(s2 + Rmin r1 r0) / 2:R
H22:(0 < S (Rlength r2))%nat -> constant_D_eq f (open_interval r r1) r3
H23:(1 < S (S (Rlength s3)))%nat -> constant_D_eq f (open_interval s2 r0) r5
open_interval r r1 x
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 = r0
r3 = r5
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:r3 = r5
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 < r0
x:=(s2 + Rmin r1 r0) / 2:R
H22:(0 < S (Rlength r2))%nat -> constant_D_eq f (open_interval r r1) r3
H23:(1 < S (S (Rlength s3)))%nat -> constant_D_eq f (open_interval s2 r0) r5

(s2 + Rmin r1 r0) / 2 < r0
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 < r0
x:=(s2 + Rmin r1 r0) / 2:R
H22:(0 < S (Rlength r2))%nat -> constant_D_eq f (open_interval r r1) r3
H23:(1 < S (S (Rlength s3)))%nat -> constant_D_eq f (open_interval s2 r0) r5
open_interval r r1 x
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 = r0
r3 = r5
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:r3 = r5
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 < r0
x:=(s2 + Rmin r1 r0) / 2:R
H22:(0 < S (Rlength r2))%nat -> constant_D_eq f (open_interval r r1) r3
H23:(1 < S (S (Rlength s3)))%nat -> constant_D_eq f (open_interval s2 r0) r5

open_interval r r1 x
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 = r0
r3 = r5
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:r3 = r5
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 < r0
x:=(s2 + Rmin r1 r0) / 2:R
H22:(0 < S (Rlength r2))%nat -> constant_D_eq f (open_interval r r1) r3
H23:(1 < S (S (Rlength s3)))%nat -> constant_D_eq f (open_interval s2 r0) r5

r < (s2 + Rmin r1 r0) / 2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 < r0
x:=(s2 + Rmin r1 r0) / 2:R
H22:(0 < S (Rlength r2))%nat -> constant_D_eq f (open_interval r r1) r3
H23:(1 < S (S (Rlength s3)))%nat -> constant_D_eq f (open_interval s2 r0) r5
(s2 + Rmin r1 r0) / 2 < r1
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 = r0
r3 = r5
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:r3 = r5
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 < r0
x:=(s2 + Rmin r1 r0) / 2:R
H22:(0 < S (Rlength r2))%nat -> constant_D_eq f (open_interval r r1) r3
H23:(1 < S (S (Rlength s3)))%nat -> constant_D_eq f (open_interval s2 r0) r5

(s2 + Rmin r1 r0) / 2 < r1
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 = r0
r3 = r5
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:r3 = r5
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:(1 < S (S (Rlength s3)))%nat -> s2 <= r0
H20:s2 <= r0
H21:s2 = r0

r3 = r5
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:r3 = r5
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1:Rlist
r0:R
s3:Rlist
r5:R
lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 (cons r5 lf2))))%nat -> pos_Rl (cons r4 (cons r5 lf2)) i <> pos_Rl (cons r4 (cons r5 lf2)) (S i) \/ f (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i)) <> pos_Rl (cons r4 (cons r5 lf2)) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> pos_Rl (cons s1 (cons s2 (cons r0 s3))) i <> pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 (cons r0 s3)))
H10:pos_Rl (cons s1 (cons s2 (cons r0 s3))) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 (cons r0 s3))) (Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3))))) = Rmax a b
H11:Rlength (cons s1 (cons s2 (cons r0 s3))) = S (Rlength (cons r4 (cons r5 lf2)))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 (cons r0 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 (cons r0 s3))) i) (pos_Rl (cons s1 (cons s2 (cons r0 s3))) (S i))) (pos_Rl (cons r4 (cons r5 lf2)) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:r < s2
Hnle:~ r1 <= s2
H16:s2 < r1
H17:r3 = r4
H18:f s2 = r3
H19:r3 = r5

r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r < r1
H15:s1 <= s2
H1:s1 = s2

r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a < b
H2:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
H3:ordered_Rlist (cons r (cons r1 r2))
H5:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H6:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H8:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H7:ordered_Rlist (cons s1 (cons s2 s3))
H10:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H9:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H11:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H13:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H12:r = s1
H14:r <= r1
H0:r = r1

r1 <= s2
rewrite <- H0; rewrite H12; apply (H7 0%nat); simpl; apply lt_O_Sn. Qed.

forall (a b : R) (f : R -> R) (l lf : Rlist), adapted_couple_opt f a b l lf -> adapted_couple_opt f b a l lf

forall (a b : R) (f : R -> R) (l lf : Rlist), adapted_couple_opt f a b l lf -> adapted_couple_opt f b a l lf
a, b:R
f:R -> R
l, lf:Rlist
H2:ordered_Rlist l
H0:pos_Rl l 0 = Rmin a b
H3:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H4:Rlength l = S (Rlength lf)
H6:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)
H5:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> pos_Rl lf i <> pos_Rl lf (S i) \/ f (pos_Rl l (S i)) <> pos_Rl lf i
H7:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> pos_Rl l i <> pos_Rl l (S i)

pos_Rl l 0 = Rmin b a
a, b:R
f:R -> R
l, lf:Rlist
H2:ordered_Rlist l
H0:pos_Rl l 0 = Rmin a b
H3:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H4:Rlength l = S (Rlength lf)
H6:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)
H5:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> pos_Rl lf i <> pos_Rl lf (S i) \/ f (pos_Rl l (S i)) <> pos_Rl lf i
H7:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> pos_Rl l i <> pos_Rl l (S i)
pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax b a
a, b:R
f:R -> R
l, lf:Rlist
H2:ordered_Rlist l
H0:pos_Rl l 0 = Rmin a b
H3:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H4:Rlength l = S (Rlength lf)
H6:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)
H5:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> pos_Rl lf i <> pos_Rl lf (S i) \/ f (pos_Rl l (S i)) <> pos_Rl lf i
H7:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> pos_Rl l i <> pos_Rl l (S i)
r:a <= b
r0:b <= a

a = b
a, b:R
f:R -> R
l, lf:Rlist
H2:ordered_Rlist l
H0:pos_Rl l 0 = Rmin a b
H3:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H4:Rlength l = S (Rlength lf)
H6:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)
H5:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> pos_Rl lf i <> pos_Rl lf (S i) \/ f (pos_Rl l (S i)) <> pos_Rl lf i
H7:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> pos_Rl l i <> pos_Rl l (S i)
n:~ a <= b
n0:~ b <= a
b = a
a, b:R
f:R -> R
l, lf:Rlist
H2:ordered_Rlist l
H0:pos_Rl l 0 = Rmin a b
H3:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H4:Rlength l = S (Rlength lf)
H6:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)
H5:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> pos_Rl lf i <> pos_Rl lf (S i) \/ f (pos_Rl l (S i)) <> pos_Rl lf i
H7:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> pos_Rl l i <> pos_Rl l (S i)
pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax b a
a, b:R
f:R -> R
l, lf:Rlist
H2:ordered_Rlist l
H0:pos_Rl l 0 = Rmin a b
H3:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H4:Rlength l = S (Rlength lf)
H6:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)
H5:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> pos_Rl lf i <> pos_Rl lf (S i) \/ f (pos_Rl l (S i)) <> pos_Rl lf i
H7:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> pos_Rl l i <> pos_Rl l (S i)
n:~ a <= b
n0:~ b <= a

b = a
a, b:R
f:R -> R
l, lf:Rlist
H2:ordered_Rlist l
H0:pos_Rl l 0 = Rmin a b
H3:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H4:Rlength l = S (Rlength lf)
H6:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)
H5:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> pos_Rl lf i <> pos_Rl lf (S i) \/ f (pos_Rl l (S i)) <> pos_Rl lf i
H7:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> pos_Rl l i <> pos_Rl l (S i)
pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax b a
a, b:R
f:R -> R
l, lf:Rlist
H2:ordered_Rlist l
H0:pos_Rl l 0 = Rmin a b
H3:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H4:Rlength l = S (Rlength lf)
H6:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)
H5:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> pos_Rl lf i <> pos_Rl lf (S i) \/ f (pos_Rl l (S i)) <> pos_Rl lf i
H7:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> pos_Rl l i <> pos_Rl l (S i)

pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax b a
a, b:R
f:R -> R
l, lf:Rlist
H2:ordered_Rlist l
H0:pos_Rl l 0 = Rmin a b
H3:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H4:Rlength l = S (Rlength lf)
H6:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)
H5:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> pos_Rl lf i <> pos_Rl lf (S i) \/ f (pos_Rl l (S i)) <> pos_Rl lf i
H7:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> pos_Rl l i <> pos_Rl l (S i)
r:a <= b
r0:b <= a

b = a
a, b:R
f:R -> R
l, lf:Rlist
H2:ordered_Rlist l
H0:pos_Rl l 0 = Rmin a b
H3:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H4:Rlength l = S (Rlength lf)
H6:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)
H5:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> pos_Rl lf i <> pos_Rl lf (S i) \/ f (pos_Rl l (S i)) <> pos_Rl lf i
H7:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> pos_Rl l i <> pos_Rl l (S i)
n:~ a <= b
n0:~ b <= a
a = b
a, b:R
f:R -> R
l, lf:Rlist
H2:ordered_Rlist l
H0:pos_Rl l 0 = Rmin a b
H3:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H4:Rlength l = S (Rlength lf)
H6:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)
H5:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> pos_Rl lf i <> pos_Rl lf (S i) \/ f (pos_Rl l (S i)) <> pos_Rl lf i
H7:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> pos_Rl l i <> pos_Rl l (S i)
n:~ a <= b
n0:~ b <= a

a = b
apply Rle_antisym; auto with real. Qed.

forall (a b r r1 r3 s1 s2 r4 : R) (r2 lf1 s3 lf2 : Rlist) (f : R -> R), a <> b -> adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1) -> adapted_couple_opt f a b (cons s1 (cons s2 s3)) (cons r4 lf2) -> r1 <= s2

forall (a b r r1 r3 s1 s2 r4 : R) (r2 lf1 s3 lf2 : Rlist) (f : R -> R), a <> b -> adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1) -> adapted_couple_opt f a b (cons s1 (cons s2 s3)) (cons r4 lf2) -> r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a <> b
H0:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H1:adapted_couple_opt f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
Hlt:a < b

r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a <> b
H0:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H1:adapted_couple_opt f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
Heq:a = b
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a <> b
H0:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H1:adapted_couple_opt f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
Hgt:a > b
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a <> b
H0:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H1:adapted_couple_opt f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
Heq:a = b

r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a <> b
H0:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H1:adapted_couple_opt f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
Hgt:a > b
r1 <= s2
a, b, r, r1, r3, s1, s2, r4:R
r2, lf1, s3, lf2:Rlist
f:R -> R
H:a <> b
H0:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H1:adapted_couple_opt f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
Hgt:a > b

r1 <= s2
eapply StepFun_P11; [ apply Hgt | apply StepFun_P2; apply H0 | apply StepFun_P12; apply H1 ]. Qed.

forall (f : R -> R) (l1 l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b l1 lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 l1 = Int_SF lf2 l2

forall (f : R -> R) (l1 l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b l1 lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 l1 = Int_SF lf2 l2
f:R -> R
l1:Rlist

forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b nil lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 nil = Int_SF lf2 l2
f:R -> R
l1:Rlist
forall (r : R) (r0 : Rlist), (forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b r0 lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 r0 = Int_SF lf2 l2) -> forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b (cons r r0) lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 (cons r r0) = Int_SF lf2 l2
f:R -> R
l1:Rlist

forall (r : R) (r0 : Rlist), (forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b r0 lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 r0 = Int_SF lf2 l2) -> forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b (cons r r0) lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 (cons r r0) = Int_SF lf2 l2
f:R -> R
l1:Rlist
r:R
r0:Rlist

(forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b nil lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 nil = Int_SF lf2 l2) -> forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b (cons r nil) lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 (cons r nil) = Int_SF lf2 l2
f:R -> R
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b r2 lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 r2 = Int_SF lf2 l2) -> forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b (cons r r2) lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 (cons r r2) = Int_SF lf2 l2) -> (forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b (cons r1 r2) lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 (cons r1 r2) = Int_SF lf2 l2) -> forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b (cons r (cons r1 r2)) lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 (cons r (cons r1 r2)) = Int_SF lf2 l2
f:R -> R
l1:Rlist
r:R
r0:Rlist
H:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 nil lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 nil = Int_SF lf3 l0
l2, lf1, lf2:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r nil) lf1
H2:adapted_couple_opt f a b l2 lf2
H3:a = b

Int_SF lf1 (cons r nil) = Int_SF lf2 l2
f:R -> R
l1:Rlist
r:R
r0:Rlist
H:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 nil lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 nil = Int_SF lf3 l0
l2, lf1, lf2:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r nil) lf1
H2:adapted_couple_opt f a b l2 lf2
H3:a <> b
Int_SF lf1 (cons r nil) = Int_SF lf2 l2
f:R -> R
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b r2 lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 r2 = Int_SF lf2 l2) -> forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b (cons r r2) lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 (cons r r2) = Int_SF lf2 l2) -> (forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b (cons r1 r2) lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 (cons r1 r2) = Int_SF lf2 l2) -> forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b (cons r (cons r1 r2)) lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 (cons r (cons r1 r2)) = Int_SF lf2 l2
f:R -> R
l1:Rlist
r:R
r0:Rlist
H:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 nil lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 nil = Int_SF lf3 l0
l2, lf1, lf2:Rlist
a, b:R
H0:a <= b
H1:adapted_couple f a b (cons r nil) lf1
H2:adapted_couple_opt f a b l2 lf2
H3:a <> b

Int_SF lf1 (cons r nil) = Int_SF lf2 l2
f:R -> R
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b r2 lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 r2 = Int_SF lf2 l2) -> forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b (cons r r2) lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 (cons r r2) = Int_SF lf2 l2) -> (forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b (cons r1 r2) lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 (cons r1 r2) = Int_SF lf2 l2) -> forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b (cons r (cons r1 r2)) lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 (cons r (cons r1 r2)) = Int_SF lf2 l2
f:R -> R
l1:Rlist
r:R
r0:Rlist

forall (r1 : R) (r2 : Rlist), ((forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b r2 lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 r2 = Int_SF lf2 l2) -> forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b (cons r r2) lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 (cons r r2) = Int_SF lf2 l2) -> (forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b (cons r1 r2) lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 (cons r1 r2) = Int_SF lf2 l2) -> forall (l2 lf1 lf2 : Rlist) (a b : R), a <= b -> adapted_couple f a b (cons r (cons r1 r2)) lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 (cons r (cons r1 r2)) = Int_SF lf2 l2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2, lf1, lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) lf1
H:adapted_couple f a b l2 lf2
H3:(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a = b

Int_SF lf1 (cons r (cons r1 r2)) = Int_SF lf2 l2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2, lf1, lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) lf1
H:adapted_couple f a b l2 lf2
H3:(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Int_SF lf1 (cons r (cons r1 r2)) = Int_SF lf2 l2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2, lf1, lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) lf1
H:adapted_couple f a b l2 lf2
H3:(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b

Int_SF lf1 (cons r (cons r1 r2)) = Int_SF lf2 l2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2, lf1, lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) lf1
H:adapted_couple f a b l2 lf2
H3:(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b

Rmin a b = a
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2, lf1, lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) lf1
H:adapted_couple f a b l2 lf2
H3:(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Int_SF lf1 (cons r (cons r1 r2)) = Int_SF lf2 l2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2, lf1, lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) lf1
H:adapted_couple f a b l2 lf2
H3:(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a

Int_SF lf1 (cons r (cons r1 r2)) = Int_SF lf2 l2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2, lf1, lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) lf1
H:adapted_couple f a b l2 lf2
H3:(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a

Rmax a b = b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2, lf1, lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) lf1
H:adapted_couple f a b l2 lf2
H3:(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
Int_SF lf1 (cons r (cons r1 r2)) = Int_SF lf2 l2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2, lf1, lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) lf1
H:adapted_couple f a b l2 lf2
H3:(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b

Int_SF lf1 (cons r (cons r1 r2)) = Int_SF lf2 l2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf1 lf0 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf1 -> adapted_couple_opt f a0 b0 l0 lf0 -> Int_SF lf1 (cons r1 r2) = Int_SF lf0 l0
l2, lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) nil
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) lf2
H3:(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)

Int_SF nil (cons r (cons r1 r2)) = Int_SF lf2 (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1, lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) lf2
H3:(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> Int_SF lf1 (cons r (cons r1 r2)) = Int_SF lf2 (cons s1 (cons s2 s3))
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF lf2 (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1, lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) lf2
H3:(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> Int_SF lf1 (cons r (cons r1 r2)) = Int_SF lf2 (cons s1 (cons s2 s3))

Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF lf2 (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf2 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf2 -> Int_SF lf0 (cons r1 r2) = Int_SF lf2 l0
l2:Rlist
r3:R
lf1:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) nil
H3:(forall i : nat, (i < Init.Nat.pred (Rlength nil))%nat -> pos_Rl nil i <> pos_Rl nil (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl nil i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)

Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF nil (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
Hreclf2:adapted_couple f a b (cons s1 (cons s2 s3)) lf2 -> (forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i)) -> Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF lf2 (cons s1 (cons s2 s3))
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
Hreclf2:adapted_couple f a b (cons s1 (cons s2 s3)) lf2 -> (forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i)) -> Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF lf2 (cons s1 (cons s2 s3))

Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)

r = s1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1

Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1

r3 = r4 \/ r = r1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r = r1

r3 = r4 \/ r = r1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
r3 = r4 \/ r = r1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1

r3 = r4 \/ r = r1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1

r1 <= s2 -> r3 = r4
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
r1 <= s2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
H8:r1 <= s2
H9:ordered_Rlist (cons s1 (cons s2 s3))
H11:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H10:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H12:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H14:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H13:ordered_Rlist (cons r (cons r1 r2))
H16:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H15:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H17:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H19:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
x:=(r + r1) / 2:R
H18:(0 < S (Rlength s3))%nat -> forall x0 : R, s1 < x0 < s2 -> f x0 = r4
H20:(0 < S (Rlength r2))%nat -> forall x0 : R, r < x0 < r1 -> f x0 = r3

r3 = f x
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
H8:r1 <= s2
H9:ordered_Rlist (cons s1 (cons s2 s3))
H11:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H10:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H12:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H14:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H13:ordered_Rlist (cons r (cons r1 r2))
H16:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H15:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H17:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H19:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
x:=(r + r1) / 2:R
H18:(0 < S (Rlength s3))%nat -> forall x0 : R, s1 < x0 < s2 -> f x0 = r4
H20:(0 < S (Rlength r2))%nat -> forall x0 : R, r < x0 < r1 -> f x0 = r3
s1 < x < s2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
r1 <= s2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
H8:r1 <= s2
H9:ordered_Rlist (cons s1 (cons s2 s3))
H11:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H10:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H12:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H14:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H13:ordered_Rlist (cons r (cons r1 r2))
H16:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H15:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H17:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H19:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
x:=(r + r1) / 2:R
H18:(0 < S (Rlength s3))%nat -> forall x0 : R, s1 < x0 < s2 -> f x0 = r4
H20:(0 < S (Rlength r2))%nat -> forall x0 : R, r < x0 < r1 -> f x0 = r3

f x = f x
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
H8:r1 <= s2
H9:ordered_Rlist (cons s1 (cons s2 s3))
H11:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H10:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H12:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H14:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H13:ordered_Rlist (cons r (cons r1 r2))
H16:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H15:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H17:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H19:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
x:=(r + r1) / 2:R
H18:(0 < S (Rlength s3))%nat -> forall x0 : R, s1 < x0 < s2 -> f x0 = r4
H20:(0 < S (Rlength r2))%nat -> forall x0 : R, r < x0 < r1 -> f x0 = r3
r < x < r1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
H8:r1 <= s2
H9:ordered_Rlist (cons s1 (cons s2 s3))
H11:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H10:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H12:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H14:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H13:ordered_Rlist (cons r (cons r1 r2))
H16:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H15:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H17:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H19:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
x:=(r + r1) / 2:R
H18:(0 < S (Rlength s3))%nat -> forall x0 : R, s1 < x0 < s2 -> f x0 = r4
H20:(0 < S (Rlength r2))%nat -> forall x0 : R, r < x0 < r1 -> f x0 = r3
s1 < x < s2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
r1 <= s2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
H8:r1 <= s2
H9:ordered_Rlist (cons s1 (cons s2 s3))
H11:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H10:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H12:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H14:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H13:ordered_Rlist (cons r (cons r1 r2))
H16:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H15:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H17:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H19:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
x:=(r + r1) / 2:R
H18:(0 < S (Rlength s3))%nat -> forall x0 : R, s1 < x0 < s2 -> f x0 = r4
H20:(0 < S (Rlength r2))%nat -> forall x0 : R, r < x0 < r1 -> f x0 = r3

r < x < r1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
H8:r1 <= s2
H9:ordered_Rlist (cons s1 (cons s2 s3))
H11:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H10:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H12:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H14:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H13:ordered_Rlist (cons r (cons r1 r2))
H16:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H15:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H17:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H19:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
x:=(r + r1) / 2:R
H18:(0 < S (Rlength s3))%nat -> forall x0 : R, s1 < x0 < s2 -> f x0 = r4
H20:(0 < S (Rlength r2))%nat -> forall x0 : R, r < x0 < r1 -> f x0 = r3
s1 < x < s2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
r1 <= s2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
H8:r1 <= s2
H9:ordered_Rlist (cons s1 (cons s2 s3))
H11:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H10:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H12:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H14:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H13:ordered_Rlist (cons r (cons r1 r2))
H16:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H15:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H17:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H19:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
x:=(r + r1) / 2:R
H18:(0 < S (Rlength s3))%nat -> forall x0 : R, s1 < x0 < s2 -> f x0 = r4
H20:(0 < S (Rlength r2))%nat -> forall x0 : R, r < x0 < r1 -> f x0 = r3
H21:r <= r1
H:r < r1

r < (r + r1) / 2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
H8:r1 <= s2
H9:ordered_Rlist (cons s1 (cons s2 s3))
H11:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H10:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H12:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H14:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H13:ordered_Rlist (cons r (cons r1 r2))
H16:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H15:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H17:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H19:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
x:=(r + r1) / 2:R
H18:(0 < S (Rlength s3))%nat -> forall x0 : R, s1 < x0 < s2 -> f x0 = r4
H20:(0 < S (Rlength r2))%nat -> forall x0 : R, r < x0 < r1 -> f x0 = r3
H21:r <= r1
H:r < r1
(r + r1) / 2 < r1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
H8:r1 <= s2
H9:ordered_Rlist (cons s1 (cons s2 s3))
H11:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H10:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H12:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H14:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H13:ordered_Rlist (cons r (cons r1 r2))
H16:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H15:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H17:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H19:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
x:=(r + r1) / 2:R
H18:(0 < S (Rlength s3))%nat -> forall x0 : R, s1 < x0 < s2 -> f x0 = r4
H20:(0 < S (Rlength r2))%nat -> forall x0 : R, r < x0 < r1 -> f x0 = r3
s1 < x < s2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
r1 <= s2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
H8:r1 <= s2
H9:ordered_Rlist (cons s1 (cons s2 s3))
H11:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H10:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H12:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H14:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H13:ordered_Rlist (cons r (cons r1 r2))
H16:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H15:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H17:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H19:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
x:=(r + r1) / 2:R
H18:(0 < S (Rlength s3))%nat -> forall x0 : R, s1 < x0 < s2 -> f x0 = r4
H20:(0 < S (Rlength r2))%nat -> forall x0 : R, r < x0 < r1 -> f x0 = r3
H21:r <= r1
H:r < r1

(r + r1) / 2 < r1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
H8:r1 <= s2
H9:ordered_Rlist (cons s1 (cons s2 s3))
H11:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H10:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H12:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H14:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H13:ordered_Rlist (cons r (cons r1 r2))
H16:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H15:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H17:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H19:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
x:=(r + r1) / 2:R
H18:(0 < S (Rlength s3))%nat -> forall x0 : R, s1 < x0 < s2 -> f x0 = r4
H20:(0 < S (Rlength r2))%nat -> forall x0 : R, r < x0 < r1 -> f x0 = r3
s1 < x < s2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
r1 <= s2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
H8:r1 <= s2
H9:ordered_Rlist (cons s1 (cons s2 s3))
H11:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H10:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H12:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H14:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H13:ordered_Rlist (cons r (cons r1 r2))
H16:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H15:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H17:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H19:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
x:=(r + r1) / 2:R
H18:(0 < S (Rlength s3))%nat -> forall x0 : R, s1 < x0 < s2 -> f x0 = r4
H20:(0 < S (Rlength r2))%nat -> forall x0 : R, r < x0 < r1 -> f x0 = r3

s1 < x < s2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
r1 <= s2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
H8:r1 <= s2
H9:ordered_Rlist (cons s1 (cons s2 s3))
H11:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H10:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H12:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H14:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H13:ordered_Rlist (cons r (cons r1 r2))
H16:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H15:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H17:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H19:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
x:=(r + r1) / 2:R
H18:(0 < S (Rlength s3))%nat -> forall x0 : R, s1 < x0 < s2 -> f x0 = r4
H20:(0 < S (Rlength r2))%nat -> forall x0 : R, r < x0 < r1 -> f x0 = r3
H21:r <= r1
H:r < r1

r < (r + r1) / 2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
H8:r1 <= s2
H9:ordered_Rlist (cons s1 (cons s2 s3))
H11:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H10:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H12:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H14:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H13:ordered_Rlist (cons r (cons r1 r2))
H16:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H15:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H17:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H19:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
x:=(r + r1) / 2:R
H18:(0 < S (Rlength s3))%nat -> forall x0 : R, s1 < x0 < s2 -> f x0 = r4
H20:(0 < S (Rlength r2))%nat -> forall x0 : R, r < x0 < r1 -> f x0 = r3
H21:r <= r1
H:r < r1
(r + r1) / 2 < s2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
r1 <= s2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
H8:r1 <= s2
H9:ordered_Rlist (cons s1 (cons s2 s3))
H11:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H10:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H12:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H14:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H13:ordered_Rlist (cons r (cons r1 r2))
H16:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H15:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H17:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H19:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
x:=(r + r1) / 2:R
H18:(0 < S (Rlength s3))%nat -> forall x0 : R, s1 < x0 < s2 -> f x0 = r4
H20:(0 < S (Rlength r2))%nat -> forall x0 : R, r < x0 < r1 -> f x0 = r3
H21:r <= r1
H:r < r1

(r + r1) / 2 < s2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
r1 <= s2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1

r1 <= s2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1

?a <> ?b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
adapted_couple ?f ?a ?b (cons ?r (cons r1 ?r2)) (cons ?r3 ?lf1)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
adapted_couple_opt ?f ?a ?b (cons ?s1 (cons s2 ?s3)) (cons ?r4 ?lf2)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1

adapted_couple ?f a b (cons ?r (cons r1 ?r2)) (cons ?r3 ?lf1)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
adapted_couple_opt ?f a b (cons ?s1 (cons s2 ?s3)) (cons ?r4 ?lf2)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1

adapted_couple_opt f a b (cons ?s1 (cons s2 ?s3)) (cons ?r4 ?lf2)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1

adapted_couple f a b (cons ?s1 (cons s2 ?s3)) (cons ?r4 ?lf2)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1
(forall i : nat, (i < Init.Nat.pred (Rlength (cons ?r4 ?lf2)))%nat -> pos_Rl (cons ?r4 ?lf2) i <> pos_Rl (cons ?r4 ?lf2) (S i) \/ f (pos_Rl (cons ?s1 (cons s2 ?s3)) (S i)) <> pos_Rl (cons ?r4 ?lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons ?s1 (cons s2 ?s3))))%nat -> pos_Rl (cons ?s1 (cons s2 ?s3)) i <> pos_Rl (cons ?s1 (cons s2 ?s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r <> r1

(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1

Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1

r1 <= s2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1

?a <> ?b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
adapted_couple ?f ?a ?b (cons ?r (cons r1 ?r2)) (cons ?r3 ?lf1)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
adapted_couple_opt ?f ?a ?b (cons ?s1 (cons s2 ?s3)) (cons ?r4 ?lf2)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1

adapted_couple ?f a b (cons ?r (cons r1 ?r2)) (cons ?r3 ?lf1)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
adapted_couple_opt ?f a b (cons ?s1 (cons s2 ?s3)) (cons ?r4 ?lf2)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1

adapted_couple_opt f a b (cons ?s1 (cons s2 ?s3)) (cons ?r4 ?lf2)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1

adapted_couple f a b (cons ?s1 (cons s2 ?s3)) (cons ?r4 ?lf2)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
(forall i : nat, (i < Init.Nat.pred (Rlength (cons ?r4 ?lf2)))%nat -> pos_Rl (cons ?r4 ?lf2) i <> pos_Rl (cons ?r4 ?lf2) (S i) \/ f (pos_Rl (cons ?s1 (cons s2 ?s3)) (S i)) <> pos_Rl (cons ?r4 ?lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons ?s1 (cons s2 ?s3))))%nat -> pos_Rl (cons ?s1 (cons s2 ?s3)) i <> pos_Rl (cons ?s1 (cons s2 ?s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1

(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2

Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4

Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2

r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2

r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r3 * (r1 - r) + r3 * (s2 - r1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2

r1 <= b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
adapted_couple f r1 b (cons r1 r2) lf1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
adapted_couple_opt f r1 b (cons r1 (cons s2 s3)) (cons r3 lf2)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)

r1 <= Rmax a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
adapted_couple f r1 b (cons r1 r2) lf1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
adapted_couple_opt f r1 b (cons r1 (cons s2 s3)) (cons r3 lf2)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2

adapted_couple f r1 b (cons r1 r2) lf1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
adapted_couple_opt f r1 b (cons r1 (cons s2 s3)) (cons r3 lf2)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2

?a <= b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
adapted_couple f ?a b (cons ?r1 (cons r1 r2)) (cons ?r3 lf1)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
adapted_couple_opt f r1 b (cons r1 (cons s2 s3)) (cons r3 lf2)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2

adapted_couple f a b (cons ?r1 (cons r1 r2)) (cons ?r3 lf1)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
adapted_couple_opt f r1 b (cons r1 (cons s2 s3)) (cons r3 lf2)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2

adapted_couple_opt f r1 b (cons r1 (cons s2 s3)) (cons r3 lf2)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2

adapted_couple f r1 b (cons r1 (cons s2 s3)) (cons r3 lf2)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2

a <= b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
adapted_couple f a b (cons a (cons r1 (cons s2 s3))) (cons r3 (cons r3 lf2))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2

adapted_couple f a b (cons a (cons r1 (cons s2 s3))) (cons r3 (cons r3 lf2))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)

r = a
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
adapted_couple f a b (cons a (cons r1 (cons s2 s3))) (cons r3 (cons r3 lf2))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a

adapted_couple f a b (cons a (cons r1 (cons s2 s3))) (cons r3 (cons r3 lf2))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a

ordered_Rlist (cons a (cons r1 (cons s2 s3)))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
pos_Rl (cons a (cons r1 (cons s2 s3))) 0 = Rmin a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
pos_Rl (cons a (cons r1 (cons s2 s3))) (Init.Nat.pred (Rlength (cons a (cons r1 (cons s2 s3))))) = Rmax a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
Rlength (cons a (cons r1 (cons s2 s3))) = S (Rlength (cons r3 (cons r3 lf2)))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons r1 (cons s2 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons a (cons r1 (cons s2 s3))) i) (pos_Rl (cons a (cons r1 (cons s2 s3))) (S i))) (pos_Rl (cons r3 (cons r3 lf2)) i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
H:(0 < S (S (Rlength s3)))%nat

pos_Rl (cons a (cons r1 (cons s2 s3))) 0 <= pos_Rl (cons a (cons r1 (cons s2 s3))) 1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S i < S (S (Rlength s3)))%nat
Hreci:(i < S (S (Rlength s3)))%nat -> pos_Rl (cons a (cons r1 (cons s2 s3))) i <= pos_Rl (cons a (cons r1 (cons s2 s3))) (S i)
pos_Rl (cons a (cons r1 (cons s2 s3))) (S i) <= pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
pos_Rl (cons a (cons r1 (cons s2 s3))) 0 = Rmin a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
pos_Rl (cons a (cons r1 (cons s2 s3))) (Init.Nat.pred (Rlength (cons a (cons r1 (cons s2 s3))))) = Rmax a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
Rlength (cons a (cons r1 (cons s2 s3))) = S (Rlength (cons r3 (cons r3 lf2)))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons r1 (cons s2 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons a (cons r1 (cons s2 s3))) i) (pos_Rl (cons a (cons r1 (cons s2 s3))) (S i))) (pos_Rl (cons r3 (cons r3 lf2)) i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
H:(0 < S (S (Rlength s3)))%nat

(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S i < S (S (Rlength s3)))%nat
Hreci:(i < S (S (Rlength s3)))%nat -> pos_Rl (cons a (cons r1 (cons s2 s3))) i <= pos_Rl (cons a (cons r1 (cons s2 s3))) (S i)
pos_Rl (cons a (cons r1 (cons s2 s3))) (S i) <= pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
pos_Rl (cons a (cons r1 (cons s2 s3))) 0 = Rmin a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
pos_Rl (cons a (cons r1 (cons s2 s3))) (Init.Nat.pred (Rlength (cons a (cons r1 (cons s2 s3))))) = Rmax a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
Rlength (cons a (cons r1 (cons s2 s3))) = S (Rlength (cons r3 (cons r3 lf2)))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons r1 (cons s2 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons a (cons r1 (cons s2 s3))) i) (pos_Rl (cons a (cons r1 (cons s2 s3))) (S i))) (pos_Rl (cons r3 (cons r3 lf2)) i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S i < S (S (Rlength s3)))%nat
Hreci:(i < S (S (Rlength s3)))%nat -> pos_Rl (cons a (cons r1 (cons s2 s3))) i <= pos_Rl (cons a (cons r1 (cons s2 s3))) (S i)

pos_Rl (cons a (cons r1 (cons s2 s3))) (S i) <= pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
pos_Rl (cons a (cons r1 (cons s2 s3))) 0 = Rmin a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
pos_Rl (cons a (cons r1 (cons s2 s3))) (Init.Nat.pred (Rlength (cons a (cons r1 (cons s2 s3))))) = Rmax a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
Rlength (cons a (cons r1 (cons s2 s3))) = S (Rlength (cons r3 (cons r3 lf2)))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons r1 (cons s2 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons a (cons r1 (cons s2 s3))) i) (pos_Rl (cons a (cons r1 (cons s2 s3))) (S i))) (pos_Rl (cons r3 (cons r3 lf2)) i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
H:(1 < S (S (Rlength s3)))%nat
Hreci:(0 < S (S (Rlength s3)))%nat -> pos_Rl (cons a (cons r1 (cons s2 s3))) 0 <= pos_Rl (cons a (cons r1 (cons s2 s3))) 1

pos_Rl (cons a (cons r1 (cons s2 s3))) 1 <= pos_Rl (cons a (cons r1 (cons s2 s3))) 2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S (S i) < S (S (Rlength s3)))%nat
Hreci:(S i < S (S (Rlength s3)))%nat -> pos_Rl (cons a (cons r1 (cons s2 s3))) (S i) <= pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i))
Hreci0:(S i < S (S (Rlength s3)))%nat -> ((i < S (S (Rlength s3)))%nat -> pos_Rl (cons a (cons r1 (cons s2 s3))) i <= pos_Rl (cons a (cons r1 (cons s2 s3))) (S i)) -> pos_Rl (cons a (cons r1 (cons s2 s3))) (S i) <= pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i))
pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i)) <= pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S (S i)))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
pos_Rl (cons a (cons r1 (cons s2 s3))) 0 = Rmin a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
pos_Rl (cons a (cons r1 (cons s2 s3))) (Init.Nat.pred (Rlength (cons a (cons r1 (cons s2 s3))))) = Rmax a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
Rlength (cons a (cons r1 (cons s2 s3))) = S (Rlength (cons r3 (cons r3 lf2)))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons r1 (cons s2 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons a (cons r1 (cons s2 s3))) i) (pos_Rl (cons a (cons r1 (cons s2 s3))) (S i))) (pos_Rl (cons r3 (cons r3 lf2)) i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S (S i) < S (S (Rlength s3)))%nat
Hreci:(S i < S (S (Rlength s3)))%nat -> pos_Rl (cons a (cons r1 (cons s2 s3))) (S i) <= pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i))
Hreci0:(S i < S (S (Rlength s3)))%nat -> ((i < S (S (Rlength s3)))%nat -> pos_Rl (cons a (cons r1 (cons s2 s3))) i <= pos_Rl (cons a (cons r1 (cons s2 s3))) (S i)) -> pos_Rl (cons a (cons r1 (cons s2 s3))) (S i) <= pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i))

pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i)) <= pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S (S i)))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
pos_Rl (cons a (cons r1 (cons s2 s3))) 0 = Rmin a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
pos_Rl (cons a (cons r1 (cons s2 s3))) (Init.Nat.pred (Rlength (cons a (cons r1 (cons s2 s3))))) = Rmax a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
Rlength (cons a (cons r1 (cons s2 s3))) = S (Rlength (cons r3 (cons r3 lf2)))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons r1 (cons s2 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons a (cons r1 (cons s2 s3))) i) (pos_Rl (cons a (cons r1 (cons s2 s3))) (S i))) (pos_Rl (cons r3 (cons r3 lf2)) i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a

pos_Rl (cons a (cons r1 (cons s2 s3))) 0 = Rmin a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
pos_Rl (cons a (cons r1 (cons s2 s3))) (Init.Nat.pred (Rlength (cons a (cons r1 (cons s2 s3))))) = Rmax a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
Rlength (cons a (cons r1 (cons s2 s3))) = S (Rlength (cons r3 (cons r3 lf2)))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons r1 (cons s2 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons a (cons r1 (cons s2 s3))) i) (pos_Rl (cons a (cons r1 (cons s2 s3))) (S i))) (pos_Rl (cons r3 (cons r3 lf2)) i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a

pos_Rl (cons a (cons r1 (cons s2 s3))) (Init.Nat.pred (Rlength (cons a (cons r1 (cons s2 s3))))) = Rmax a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
Rlength (cons a (cons r1 (cons s2 s3))) = S (Rlength (cons r3 (cons r3 lf2)))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons r1 (cons s2 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons a (cons r1 (cons s2 s3))) i) (pos_Rl (cons a (cons r1 (cons s2 s3))) (S i))) (pos_Rl (cons r3 (cons r3 lf2)) i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a

Rlength (cons a (cons r1 (cons s2 s3))) = S (Rlength (cons r3 (cons r3 lf2)))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons r1 (cons s2 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons a (cons r1 (cons s2 s3))) i) (pos_Rl (cons a (cons r1 (cons s2 s3))) (S i))) (pos_Rl (cons r3 (cons r3 lf2)) i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a

forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons r1 (cons s2 s3)))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons a (cons r1 (cons s2 s3))) i) (pos_Rl (cons a (cons r1 (cons s2 s3))) (S i))) (pos_Rl (cons r3 (cons r3 lf2)) i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
H:(0 < S (S (Rlength s3)))%nat
x:R
H2:pos_Rl (cons a (cons r1 (cons s2 s3))) 0 < x < pos_Rl (cons a (cons r1 (cons s2 s3))) 1

f x = pos_Rl (cons r3 (cons r3 lf2)) 0
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S i < S (S (Rlength s3)))%nat
x:R
H2:pos_Rl (cons a (cons r1 (cons s2 s3))) (S i) < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i))
Hreci:(i < S (S (Rlength s3)))%nat -> pos_Rl (cons a (cons r1 (cons s2 s3))) i < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S i) -> f x = pos_Rl (cons r3 (cons r3 lf2)) i
f x = pos_Rl (cons r3 (cons r3 lf2)) (S i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
H:(0 < S (S (Rlength s3)))%nat
x:R
H2:pos_Rl (cons a (cons r1 (cons s2 s3))) 0 < x < pos_Rl (cons a (cons r1 (cons s2 s3))) 1

(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
H:(0 < S (S (Rlength s3)))%nat
x:R
H2:pos_Rl (cons a (cons r1 (cons s2 s3))) 0 < x < pos_Rl (cons a (cons r1 (cons s2 s3))) 1
open_interval (pos_Rl (cons r (cons r1 r2)) 0) (pos_Rl (cons r (cons r1 r2)) 1) x
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S i < S (S (Rlength s3)))%nat
x:R
H2:pos_Rl (cons a (cons r1 (cons s2 s3))) (S i) < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i))
Hreci:(i < S (S (Rlength s3)))%nat -> pos_Rl (cons a (cons r1 (cons s2 s3))) i < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S i) -> f x = pos_Rl (cons r3 (cons r3 lf2)) i
f x = pos_Rl (cons r3 (cons r3 lf2)) (S i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
H:(0 < S (S (Rlength s3)))%nat
x:R
H2:pos_Rl (cons a (cons r1 (cons s2 s3))) 0 < x < pos_Rl (cons a (cons r1 (cons s2 s3))) 1

open_interval (pos_Rl (cons r (cons r1 r2)) 0) (pos_Rl (cons r (cons r1 r2)) 1) x
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S i < S (S (Rlength s3)))%nat
x:R
H2:pos_Rl (cons a (cons r1 (cons s2 s3))) (S i) < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i))
Hreci:(i < S (S (Rlength s3)))%nat -> pos_Rl (cons a (cons r1 (cons s2 s3))) i < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S i) -> f x = pos_Rl (cons r3 (cons r3 lf2)) i
f x = pos_Rl (cons r3 (cons r3 lf2)) (S i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S i < S (S (Rlength s3)))%nat
x:R
H2:pos_Rl (cons a (cons r1 (cons s2 s3))) (S i) < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i))
Hreci:(i < S (S (Rlength s3)))%nat -> pos_Rl (cons a (cons r1 (cons s2 s3))) i < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S i) -> f x = pos_Rl (cons r3 (cons r3 lf2)) i

f x = pos_Rl (cons r3 (cons r3 lf2)) (S i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
H:(1 < S (S (Rlength s3)))%nat
x:R
H2:pos_Rl (cons a (cons r1 (cons s2 s3))) 1 < x < pos_Rl (cons a (cons r1 (cons s2 s3))) 2

f x = pos_Rl (cons r3 (cons r3 lf2)) 1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S (S i) < S (S (Rlength s3)))%nat
x:R
H2:pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i)) < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S (S i)))
Hreci:(S i < S (S (Rlength s3)))%nat -> pos_Rl (cons a (cons r1 (cons s2 s3))) (S i) < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i)) -> f x = pos_Rl (cons r3 (cons r3 lf2)) (S i)
f x = pos_Rl (cons r3 (cons r3 lf2)) (S (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
H:(1 < S (S (Rlength s3)))%nat
x:R
H2:r1 < x < s2

(0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
H:(1 < S (S (Rlength s3)))%nat
x:R
H2:r1 < x < s2
open_interval (pos_Rl (cons s1 (cons s2 s3)) 0) (pos_Rl (cons s1 (cons s2 s3)) 1) x
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S (S i) < S (S (Rlength s3)))%nat
x:R
H2:pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i)) < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S (S i)))
Hreci:(S i < S (S (Rlength s3)))%nat -> pos_Rl (cons a (cons r1 (cons s2 s3))) (S i) < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i)) -> f x = pos_Rl (cons r3 (cons r3 lf2)) (S i)
f x = pos_Rl (cons r3 (cons r3 lf2)) (S (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
H:(1 < S (S (Rlength s3)))%nat
x:R
H2:r1 < x < s2

open_interval (pos_Rl (cons s1 (cons s2 s3)) 0) (pos_Rl (cons s1 (cons s2 s3)) 1) x
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S (S i) < S (S (Rlength s3)))%nat
x:R
H2:pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i)) < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S (S i)))
Hreci:(S i < S (S (Rlength s3)))%nat -> pos_Rl (cons a (cons r1 (cons s2 s3))) (S i) < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i)) -> f x = pos_Rl (cons r3 (cons r3 lf2)) (S i)
f x = pos_Rl (cons r3 (cons r3 lf2)) (S (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
H:(1 < S (S (Rlength s3)))%nat
x:R
H2:r1 < x < s2
H22:r1 < x
H23:x < s2

s1 < x
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
H:(1 < S (S (Rlength s3)))%nat
x:R
H2:r1 < x < s2
H22:r1 < x
H23:x < s2
x < s2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S (S i) < S (S (Rlength s3)))%nat
x:R
H2:pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i)) < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S (S i)))
Hreci:(S i < S (S (Rlength s3)))%nat -> pos_Rl (cons a (cons r1 (cons s2 s3))) (S i) < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i)) -> f x = pos_Rl (cons r3 (cons r3 lf2)) (S i)
f x = pos_Rl (cons r3 (cons r3 lf2)) (S (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
H:(1 < S (S (Rlength s3)))%nat
x:R
H2:r1 < x < s2
H22:r1 < x
H23:x < s2

x < s2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S (S i) < S (S (Rlength s3)))%nat
x:R
H2:pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i)) < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S (S i)))
Hreci:(S i < S (S (Rlength s3)))%nat -> pos_Rl (cons a (cons r1 (cons s2 s3))) (S i) < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i)) -> f x = pos_Rl (cons r3 (cons r3 lf2)) (S i)
f x = pos_Rl (cons r3 (cons r3 lf2)) (S (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S (S i) < S (S (Rlength s3)))%nat
x:R
H2:pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i)) < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S (S i)))
Hreci:(S i < S (S (Rlength s3)))%nat -> pos_Rl (cons a (cons r1 (cons s2 s3))) (S i) < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i)) -> f x = pos_Rl (cons r3 (cons r3 lf2)) (S i)

f x = pos_Rl (cons r3 (cons r3 lf2)) (S (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S (S i) < S (S (Rlength s3)))%nat
x:R
H2:pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i)) < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S (S i)))

(S i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S (S i) < S (S (Rlength s3)))%nat
x:R
H2:pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i)) < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S (S i)))
open_interval (pos_Rl (cons s1 (cons s2 s3)) (S i)) (pos_Rl (cons s1 (cons s2 s3)) (S (S i))) x
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S (S i) < S (S (Rlength s3)))%nat
x:R
H2:pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S i)) < x < pos_Rl (cons a (cons r1 (cons s2 s3))) (S (S (S i)))

open_interval (pos_Rl (cons s1 (cons s2 s3)) (S i)) (pos_Rl (cons s1 (cons s2 s3)) (S (S i))) x
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2

(forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H3:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i
H11:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i)

forall i : nat, (i < Init.Nat.pred (Rlength (cons r3 lf2)))%nat -> pos_Rl (cons r3 lf2) i <> pos_Rl (cons r3 lf2) (S i) \/ f (pos_Rl (cons r1 (cons s2 s3)) (S i)) <> pos_Rl (cons r3 lf2) i
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H3:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i
H11:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H3:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i
H11:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i)

forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 (cons s2 s3))))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H3:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i
H11:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i <> pos_Rl (cons s1 (cons s2 s3)) (S i)
H12:(0 < S (Rlength s3))%nat

pos_Rl (cons r1 (cons s2 s3)) 0 <> pos_Rl (cons r1 (cons s2 s3)) 1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H3:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0
H11:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i0 <> pos_Rl (cons s1 (cons s2 s3)) (S i0)
i:nat
H12:(S i < S (Rlength s3))%nat
Hreci:(i < S (Rlength s3))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i)
pos_Rl (cons r1 (cons s2 s3)) (S i) <> pos_Rl (cons r1 (cons s2 s3)) (S (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 < s2
H3:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0
H11:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> pos_Rl (cons s1 (cons s2 s3)) i0 <> pos_Rl (cons s1 (cons s2 s3)) (S i0)
i:nat
H12:(S i < S (Rlength s3))%nat
Hreci:(i < S (Rlength s3))%nat -> pos_Rl (cons r1 (cons s2 s3)) i <> pos_Rl (cons r1 (cons s2 s3)) (S i)

pos_Rl (cons r1 (cons s2 s3)) (S i) <> pos_Rl (cons r1 (cons s2 s3)) (S (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2

r3 * (r1 - r) + Int_SF lf1 (cons r1 r2) = r4 * (s2 - s1) + Int_SF lf2 (cons s2 s3)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2

r1 <= b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
adapted_couple f r1 b (cons r1 r2) lf1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
adapted_couple_opt f r1 b (cons r1 s3) lf2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)

r1 <= Rmax a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
adapted_couple f r1 b (cons r1 r2) lf1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
adapted_couple_opt f r1 b (cons r1 s3) lf2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2

adapted_couple f r1 b (cons r1 r2) lf1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
adapted_couple_opt f r1 b (cons r1 s3) lf2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2

?a <= b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
adapted_couple f ?a b (cons ?r1 (cons r1 r2)) (cons ?r3 lf1)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
adapted_couple_opt f r1 b (cons r1 s3) lf2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2

adapted_couple f a b (cons ?r1 (cons r1 r2)) (cons ?r3 lf1)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
adapted_couple_opt f r1 b (cons r1 s3) lf2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2

adapted_couple_opt f r1 b (cons r1 s3) lf2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2

adapted_couple f r1 b (cons r1 s3) lf2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl (cons r1 s3) (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2

a <= b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
adapted_couple f a b (cons a (cons r1 s3)) (cons r3 lf2)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl (cons r1 s3) (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2

adapted_couple f a b (cons a (cons r1 s3)) (cons r3 lf2)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl (cons r1 s3) (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)

r = a
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
adapted_couple f a b (cons a (cons r1 s3)) (cons r3 lf2)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl (cons r1 s3) (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a

adapted_couple f a b (cons a (cons r1 s3)) (cons r3 lf2)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl (cons r1 s3) (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a

ordered_Rlist (cons a (cons r1 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
pos_Rl (cons a (cons r1 s3)) 0 = Rmin a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
pos_Rl (cons a (cons r1 s3)) (Init.Nat.pred (Rlength (cons a (cons r1 s3)))) = Rmax a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
Rlength (cons a (cons r1 s3)) = S (Rlength (cons r3 lf2))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons r1 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons a (cons r1 s3)) i) (pos_Rl (cons a (cons r1 s3)) (S i))) (pos_Rl (cons r3 lf2) i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl (cons r1 s3) (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
H:(0 < S (Rlength s3))%nat

pos_Rl (cons a (cons r1 s3)) 0 <= pos_Rl (cons a (cons r1 s3)) 1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S i < S (Rlength s3))%nat
Hreci:(i < S (Rlength s3))%nat -> pos_Rl (cons a (cons r1 s3)) i <= pos_Rl (cons a (cons r1 s3)) (S i)
pos_Rl (cons a (cons r1 s3)) (S i) <= pos_Rl (cons a (cons r1 s3)) (S (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
pos_Rl (cons a (cons r1 s3)) 0 = Rmin a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
pos_Rl (cons a (cons r1 s3)) (Init.Nat.pred (Rlength (cons a (cons r1 s3)))) = Rmax a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
Rlength (cons a (cons r1 s3)) = S (Rlength (cons r3 lf2))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons r1 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons a (cons r1 s3)) i) (pos_Rl (cons a (cons r1 s3)) (S i))) (pos_Rl (cons r3 lf2) i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl (cons r1 s3) (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S i < S (Rlength s3))%nat
Hreci:(i < S (Rlength s3))%nat -> pos_Rl (cons a (cons r1 s3)) i <= pos_Rl (cons a (cons r1 s3)) (S i)

pos_Rl (cons a (cons r1 s3)) (S i) <= pos_Rl (cons a (cons r1 s3)) (S (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
pos_Rl (cons a (cons r1 s3)) 0 = Rmin a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
pos_Rl (cons a (cons r1 s3)) (Init.Nat.pred (Rlength (cons a (cons r1 s3)))) = Rmax a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
Rlength (cons a (cons r1 s3)) = S (Rlength (cons r3 lf2))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons r1 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons a (cons r1 s3)) i) (pos_Rl (cons a (cons r1 s3)) (S i))) (pos_Rl (cons r3 lf2) i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl (cons r1 s3) (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a

pos_Rl (cons a (cons r1 s3)) 0 = Rmin a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
pos_Rl (cons a (cons r1 s3)) (Init.Nat.pred (Rlength (cons a (cons r1 s3)))) = Rmax a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
Rlength (cons a (cons r1 s3)) = S (Rlength (cons r3 lf2))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons r1 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons a (cons r1 s3)) i) (pos_Rl (cons a (cons r1 s3)) (S i))) (pos_Rl (cons r3 lf2) i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl (cons r1 s3) (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a

pos_Rl (cons a (cons r1 s3)) (Init.Nat.pred (Rlength (cons a (cons r1 s3)))) = Rmax a b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
Rlength (cons a (cons r1 s3)) = S (Rlength (cons r3 lf2))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons r1 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons a (cons r1 s3)) i) (pos_Rl (cons a (cons r1 s3)) (S i))) (pos_Rl (cons r3 lf2) i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl (cons r1 s3) (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a

Rlength (cons a (cons r1 s3)) = S (Rlength (cons r3 lf2))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons r1 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons a (cons r1 s3)) i) (pos_Rl (cons a (cons r1 s3)) (S i))) (pos_Rl (cons r3 lf2) i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl (cons r1 s3) (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a

forall i : nat, (i < Init.Nat.pred (Rlength (cons a (cons r1 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons a (cons r1 s3)) i) (pos_Rl (cons a (cons r1 s3)) (S i))) (pos_Rl (cons r3 lf2) i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl (cons r1 s3) (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
H:(0 < S (Rlength s3))%nat
x:R
H2:pos_Rl (cons a (cons r1 s3)) 0 < x < pos_Rl (cons a (cons r1 s3)) 1

f x = pos_Rl (cons r3 lf2) 0
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S i < S (Rlength s3))%nat
x:R
H2:pos_Rl (cons a (cons r1 s3)) (S i) < x < pos_Rl (cons a (cons r1 s3)) (S (S i))
Hreci:(i < S (Rlength s3))%nat -> pos_Rl (cons a (cons r1 s3)) i < x < pos_Rl (cons a (cons r1 s3)) (S i) -> f x = pos_Rl (cons r3 lf2) i
f x = pos_Rl (cons r3 lf2) (S i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl (cons r1 s3) (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
H:(0 < S (Rlength s3))%nat
x:R
H2:pos_Rl (cons a (cons r1 s3)) 0 < x < pos_Rl (cons a (cons r1 s3)) 1

(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
H:(0 < S (Rlength s3))%nat
x:R
H2:pos_Rl (cons a (cons r1 s3)) 0 < x < pos_Rl (cons a (cons r1 s3)) 1
open_interval (pos_Rl (cons r (cons r1 r2)) 0) (pos_Rl (cons r (cons r1 r2)) 1) x
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S i < S (Rlength s3))%nat
x:R
H2:pos_Rl (cons a (cons r1 s3)) (S i) < x < pos_Rl (cons a (cons r1 s3)) (S (S i))
Hreci:(i < S (Rlength s3))%nat -> pos_Rl (cons a (cons r1 s3)) i < x < pos_Rl (cons a (cons r1 s3)) (S i) -> f x = pos_Rl (cons r3 lf2) i
f x = pos_Rl (cons r3 lf2) (S i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl (cons r1 s3) (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i : nat, (i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i) (pos_Rl (cons s1 (cons s2 s3)) (S i))) (pos_Rl (cons r4 lf2) i)
H20:r = a
H:(0 < S (Rlength s3))%nat
x:R
H2:pos_Rl (cons a (cons r1 s3)) 0 < x < pos_Rl (cons a (cons r1 s3)) 1

open_interval (pos_Rl (cons r (cons r1 r2)) 0) (pos_Rl (cons r (cons r1 r2)) 1) x
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S i < S (Rlength s3))%nat
x:R
H2:pos_Rl (cons a (cons r1 s3)) (S i) < x < pos_Rl (cons a (cons r1 s3)) (S (S i))
Hreci:(i < S (Rlength s3))%nat -> pos_Rl (cons a (cons r1 s3)) i < x < pos_Rl (cons a (cons r1 s3)) (S i) -> f x = pos_Rl (cons r3 lf2) i
f x = pos_Rl (cons r3 lf2) (S i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl (cons r1 s3) (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S i < S (Rlength s3))%nat
x:R
H2:pos_Rl (cons a (cons r1 s3)) (S i) < x < pos_Rl (cons a (cons r1 s3)) (S (S i))
Hreci:(i < S (Rlength s3))%nat -> pos_Rl (cons a (cons r1 s3)) i < x < pos_Rl (cons a (cons r1 s3)) (S i) -> f x = pos_Rl (cons r3 lf2) i

f x = pos_Rl (cons r3 lf2) (S i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl (cons r1 s3) (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S i < S (Rlength s3))%nat
x:R
H2:pos_Rl (cons a (cons r1 s3)) (S i) < x < pos_Rl (cons a (cons r1 s3)) (S (S i))

(S i < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S i < S (Rlength s3))%nat
x:R
H2:pos_Rl (cons a (cons r1 s3)) (S i) < x < pos_Rl (cons a (cons r1 s3)) (S (S i))
open_interval (pos_Rl (cons s1 (cons s2 s3)) (S i)) (pos_Rl (cons s1 (cons s2 s3)) (S (S i))) x
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl (cons r1 s3) (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
s1, s2:R
s3:Rlist
H3:(forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl l2 (S i0)) <> pos_Rl (cons r4 lf2) i0) /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H11:ordered_Rlist (cons r (cons r1 r2))
H13:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H12:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H14:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H16:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H15:ordered_Rlist (cons s1 (cons s2 s3))
H18:pos_Rl (cons s1 (cons s2 s3)) 0 = Rmin a b
H17:pos_Rl (cons s1 (cons s2 s3)) (Init.Nat.pred (Rlength (cons s1 (cons s2 s3)))) = Rmax a b
H19:Rlength (cons s1 (cons s2 s3)) = S (Rlength (cons r4 lf2))
H21:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons s1 (cons s2 s3))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons s1 (cons s2 s3)) i0) (pos_Rl (cons s1 (cons s2 s3)) (S i0))) (pos_Rl (cons r4 lf2) i0)
H20:r = a
i:nat
H:(S i < S (Rlength s3))%nat
x:R
H2:pos_Rl (cons a (cons r1 s3)) (S i) < x < pos_Rl (cons a (cons r1 s3)) (S (S i))

open_interval (pos_Rl (cons s1 (cons s2 s3)) (S i)) (pos_Rl (cons s1 (cons s2 s3)) (S (S i))) x
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl (cons r1 s3) (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2

(forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl (cons r1 s3) (S i)) <> pos_Rl lf2 i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H3:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i
H11:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i)

forall i : nat, (i < Init.Nat.pred (Rlength lf2))%nat -> pos_Rl lf2 i <> pos_Rl lf2 (S i) \/ f (pos_Rl (cons r1 s3) (S i)) <> pos_Rl lf2 i
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H3:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i
H11:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H3:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i0)) <> pos_Rl (cons r4 lf2) i0
H11:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0)
i:nat
H12:(i < Init.Nat.pred (Rlength lf2))%nat

(S i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H3:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i
H11:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H3:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i0)) <> pos_Rl (cons r4 lf2) i0
H11:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0)
i:nat
H12:(i < Init.Nat.pred (Rlength lf2))%nat

(S i < S (Init.Nat.pred (Rlength lf2)))%nat
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H3:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i0)) <> pos_Rl (cons r4 lf2) i0
H11:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0)
i:nat
H12:(i < Init.Nat.pred (Rlength lf2))%nat
S (Init.Nat.pred (Rlength lf2)) = Rlength lf2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H3:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i
H11:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H3:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i0 <> pos_Rl (cons r4 lf2) (S i0) \/ f (pos_Rl (cons s1 (cons s2 s3)) (S i0)) <> pos_Rl (cons r4 lf2) i0
H11:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i0 <> pos_Rl l2 (S i0)
i:nat
H12:(i < Init.Nat.pred (Rlength lf2))%nat

S (Init.Nat.pred (Rlength lf2)) = Rlength lf2
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H3:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i
H11:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r3 = r4
H10:r1 = s2
H3:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i
H11:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i)

forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 s3)))%nat -> pos_Rl (cons r1 s3) i <> pos_Rl (cons r1 s3) (S i)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1

Int_SF (cons r3 lf1) (cons r (cons r1 r2)) = Int_SF (cons r4 lf2) (cons s1 (cons s2 s3))
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1

?a <= ?b
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
adapted_couple f ?a ?b (cons r1 r2) lf1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
adapted_couple_opt f ?a ?b (cons s1 (cons s2 s3)) (cons r4 lf2)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1

adapted_couple f a b (cons r1 r2) lf1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
adapted_couple_opt f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1

adapted_couple f a b (cons r1 r2) lf1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1

r = a
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
H10:r = a
adapted_couple f a b (cons r1 r2) lf1
f:R -> R
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
H0:forall (l0 lf0 lf3 : Rlist) (a0 b0 : R), a0 <= b0 -> adapted_couple f a0 b0 (cons r1 r2) lf0 -> adapted_couple_opt f a0 b0 l0 lf3 -> Int_SF lf0 (cons r1 r2) = Int_SF lf3 l0
l2:Rlist
r3:R
lf1:Rlist
r4:R
lf2:Rlist
a, b:R
H1:a <= b
H2:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
s1, s2:R
s3:Rlist
H:adapted_couple f a b (cons s1 (cons s2 s3)) (cons r4 lf2)
H3:(forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 lf2)))%nat -> pos_Rl (cons r4 lf2) i <> pos_Rl (cons r4 lf2) (S i) \/ f (pos_Rl l2 (S i)) <> pos_Rl (cons r4 lf2) i) /\ (forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> pos_Rl l2 i <> pos_Rl l2 (S i))
H4:a <> b
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H5:l2 = cons s1 (cons s2 s3)
H6:r = s1
H7:r3 = r4 \/ r = r1
H8:r1 <= s2
H9:r = r1
H10:r = a

adapted_couple f a b (cons r1 r2) lf1
rewrite <- H9; rewrite H10; apply StepFun_P7 with a r r3; [ apply H1 | pattern a at 2; rewrite <- H10; pattern r at 2; rewrite H9; apply H2 ]. Qed.

forall (f : R -> R) (l1 l2 lf1 lf2 : Rlist) (a b : R), adapted_couple f a b l1 lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 l1 = Int_SF lf2 l2

forall (f : R -> R) (l1 l2 lf1 lf2 : Rlist) (a b : R), adapted_couple f a b l1 lf1 -> adapted_couple_opt f a b l2 lf2 -> Int_SF lf1 l1 = Int_SF lf2 l2
intros; destruct (Rle_dec a b) as [Hle|Hnle]; [ apply (StepFun_P14 Hle H H0) | assert (H1 : b <= a); [ auto with real | eapply StepFun_P14; [ apply H1 | apply StepFun_P2; apply H | apply StepFun_P12; apply H0 ] ] ]. Qed.

forall (f : R -> R) (l lf : Rlist) (a b : R), adapted_couple f a b l lf -> exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'

forall (f : R -> R) (l lf : Rlist) (a b : R), adapted_couple f a b l lf -> exists l' lf' : Rlist, adapted_couple_opt f a b l' lf'
intros; destruct (Rle_dec a b) as [Hle|Hnle]; [ apply (StepFun_P10 Hle H) | assert (H1 : b <= a); [ auto with real | assert (H2 := StepFun_P10 H1 (StepFun_P2 H)); elim H2; intros l' [lf' H3]; exists l'; exists lf'; apply StepFun_P12; assumption ] ]. Qed.

forall (f : R -> R) (l1 l2 lf1 lf2 : Rlist) (a b : R), adapted_couple f a b l1 lf1 -> adapted_couple f a b l2 lf2 -> Int_SF lf1 l1 = Int_SF lf2 l2

forall (f : R -> R) (l1 l2 lf1 lf2 : Rlist) (a b : R), adapted_couple f a b l1 lf1 -> adapted_couple f a b l2 lf2 -> Int_SF lf1 l1 = Int_SF lf2 l2
intros; elim (StepFun_P16 H); intros l' [lf' H1]; rewrite (StepFun_P15 H H1); rewrite (StepFun_P15 H0 H1); reflexivity. Qed.

forall a b c : R, RiemannInt_SF {| fe := fct_cte c; pre := StepFun_P4 a b c |} = c * (b - a)

forall a b c : R, RiemannInt_SF {| fe := fct_cte c; pre := StepFun_P4 a b c |} = c * (b - a)
a, b, c:R
r:a <= b

Int_SF (subdivision_val {| fe := fct_cte c; pre := StepFun_P4 a b c |}) (subdivision {| fe := fct_cte c; pre := StepFun_P4 a b c |}) = c * (b - a)
a, b, c:R
n:~ a <= b
- Int_SF (subdivision_val {| fe := fct_cte c; pre := StepFun_P4 a b c |}) (subdivision {| fe := fct_cte c; pre := StepFun_P4 a b c |}) = c * (b - a)
a, b, c:R
n:~ a <= b

- Int_SF (subdivision_val {| fe := fct_cte c; pre := StepFun_P4 a b c |}) (subdivision {| fe := fct_cte c; pre := StepFun_P4 a b c |}) = c * (b - a)
replace (Int_SF (subdivision_val (mkStepFun (StepFun_P4 a b c))) (subdivision (mkStepFun (StepFun_P4 a b c)))) with (Int_SF (cons c nil) (cons b (cons a nil))); [ simpl; ring | apply StepFun_P17 with (fct_cte c) a b; [ apply StepFun_P2; apply StepFun_P3; auto with real | apply (StepFun_P1 (mkStepFun (StepFun_P4 a b c))) ] ]. Qed.

forall (l1 : Rlist) (f g : R -> R) (l : R), Int_SF (FF l1 (fun x : R => f x + l * g x)) l1 = Int_SF (FF l1 f) l1 + l * Int_SF (FF l1 g) l1

forall (l1 : Rlist) (f g : R -> R) (l : R), Int_SF (FF l1 (fun x : R => f x + l * g x)) l1 = Int_SF (FF l1 f) l1 + l * Int_SF (FF l1 g) l1
intros; induction l1 as [| r l1 Hrecl1]; [ simpl; ring | induction l1 as [| r0 l1 Hrecl0]; simpl; [ ring | simpl in Hrecl1; rewrite Hrecl1; ring ] ]. Qed.

forall (l : Rlist) (f : R -> R), (0 < Rlength l)%nat -> Rlength l = S (Rlength (FF l f))

forall (l : Rlist) (f : R -> R), (0 < Rlength l)%nat -> Rlength l = S (Rlength (FF l f))
intros l f H; induction l; [ elim (lt_irrefl _ H) | simpl; rewrite RList_P18; rewrite RList_P14; reflexivity ]. Qed.

forall (a b : R) (f : R -> R) (l : Rlist), is_subdivision f a b l -> adapted_couple f a b l (FF l f)

forall (a b : R) (f : R -> R) (l : Rlist), is_subdivision f a b l -> adapted_couple f a b l (FF l f)
a, b:R
f:R -> R
l, x:Rlist
H:ordered_Rlist l
H1:pos_Rl l 0 = Rmin a b
H0:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H2:Rlength l = S (Rlength x)
H4:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl x i)

adapted_couple f a b l (FF l f)
a, b:R
f:R -> R
l, x:Rlist
H:ordered_Rlist l
H1:pos_Rl l 0 = Rmin a b
H0:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H2:Rlength l = S (Rlength x)
H4:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl x i)

Rlength l = S (Rlength (FF l f))
a, b:R
f:R -> R
l, x:Rlist
H:ordered_Rlist l
H1:pos_Rl l 0 = Rmin a b
H0:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H2:Rlength l = S (Rlength x)
H4:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl x i)
forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl (FF l f) i)
a, b:R
f:R -> R
l, x:Rlist
H:ordered_Rlist l
H1:pos_Rl l 0 = Rmin a b
H0:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H2:Rlength l = S (Rlength x)
H4:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl x i)

forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl (FF l f) i)
a, b:R
f:R -> R
x:Rlist
H:ordered_Rlist nil
H1:pos_Rl nil 0 = Rmin a b
H0:pos_Rl nil (Init.Nat.pred (Rlength nil)) = Rmax a b
H2:Rlength nil = S (Rlength x)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength nil))%nat -> constant_D_eq f (open_interval (pos_Rl nil i0) (pos_Rl nil (S i0))) (pos_Rl x i0)
i:nat
H3:(i < Init.Nat.pred (Rlength nil))%nat
H5:forall x1 : R, pos_Rl nil i < x1 < pos_Rl nil (S i) -> f x1 = pos_Rl x i
x0:R
H6:pos_Rl nil i < x0 < pos_Rl nil (S i)

f x0 = pos_Rl (FF nil f) i
a, b:R
f:R -> R
r:R
l, x:Rlist
H:ordered_Rlist (cons r l)
H1:pos_Rl (cons r l) 0 = Rmin a b
H0:pos_Rl (cons r l) (Init.Nat.pred (Rlength (cons r l))) = Rmax a b
H2:Rlength (cons r l) = S (Rlength x)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r l) i0) (pos_Rl (cons r l) (S i0))) (pos_Rl x i0)
i:nat
H3:(i < Init.Nat.pred (Rlength (cons r l)))%nat
H5:forall x1 : R, pos_Rl (cons r l) i < x1 < pos_Rl (cons r l) (S i) -> f x1 = pos_Rl x i
x0:R
H6:pos_Rl (cons r l) i < x0 < pos_Rl (cons r l) (S i)
Hrecl:ordered_Rlist l -> pos_Rl l 0 = Rmin a b -> pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b -> Rlength l = S (Rlength x) -> (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i0) (pos_Rl l (S i0))) (pos_Rl x i0)) -> (i < Init.Nat.pred (Rlength l))%nat -> (forall x1 : R, pos_Rl l i < x1 < pos_Rl l (S i) -> f x1 = pos_Rl x i) -> pos_Rl l i < x0 < pos_Rl l (S i) -> f x0 = pos_Rl (FF l f) i
f x0 = pos_Rl (FF (cons r l) f) i
a, b:R
f:R -> R
r:R
l, x:Rlist
H:ordered_Rlist (cons r l)
H1:pos_Rl (cons r l) 0 = Rmin a b
H0:pos_Rl (cons r l) (Init.Nat.pred (Rlength (cons r l))) = Rmax a b
H2:Rlength (cons r l) = S (Rlength x)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r l) i0) (pos_Rl (cons r l) (S i0))) (pos_Rl x i0)
i:nat
H3:(i < Init.Nat.pred (Rlength (cons r l)))%nat
H5:forall x1 : R, pos_Rl (cons r l) i < x1 < pos_Rl (cons r l) (S i) -> f x1 = pos_Rl x i
x0:R
H6:pos_Rl (cons r l) i < x0 < pos_Rl (cons r l) (S i)
Hrecl:ordered_Rlist l -> pos_Rl l 0 = Rmin a b -> pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b -> Rlength l = S (Rlength x) -> (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i0) (pos_Rl l (S i0))) (pos_Rl x i0)) -> (i < Init.Nat.pred (Rlength l))%nat -> (forall x1 : R, pos_Rl l i < x1 < pos_Rl l (S i) -> f x1 = pos_Rl x i) -> pos_Rl l i < x0 < pos_Rl l (S i) -> f x0 = pos_Rl (FF l f) i

f x0 = pos_Rl (FF (cons r l) f) i
a, b:R
f:R -> R
r:R
l, x:Rlist
H:ordered_Rlist (cons r l)
H1:pos_Rl (cons r l) 0 = Rmin a b
H0:pos_Rl (cons r l) (Init.Nat.pred (Rlength (cons r l))) = Rmax a b
H2:Rlength (cons r l) = S (Rlength x)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r l) i0) (pos_Rl (cons r l) (S i0))) (pos_Rl x i0)
i:nat
H3:(i < Init.Nat.pred (Rlength (cons r l)))%nat
H5:forall x1 : R, pos_Rl (cons r l) i < x1 < pos_Rl (cons r l) (S i) -> f x1 = pos_Rl x i
x0:R
H6:pos_Rl (cons r l) i < x0 < pos_Rl (cons r l) (S i)
Hrecl:ordered_Rlist l -> pos_Rl l 0 = Rmin a b -> pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b -> Rlength l = S (Rlength x) -> (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i0) (pos_Rl l (S i0))) (pos_Rl x i0)) -> (i < Init.Nat.pred (Rlength l))%nat -> (forall x1 : R, pos_Rl l i < x1 < pos_Rl l (S i) -> f x1 = pos_Rl x i) -> pos_Rl l i < x0 < pos_Rl l (S i) -> f x0 = pos_Rl (FF l f) i

f x0 = f (pos_Rl (mid_Rlist l r) i)
a, b:R
f:R -> R
r:R
l, x:Rlist
H:ordered_Rlist (cons r l)
H1:pos_Rl (cons r l) 0 = Rmin a b
H0:pos_Rl (cons r l) (Init.Nat.pred (Rlength (cons r l))) = Rmax a b
H2:Rlength (cons r l) = S (Rlength x)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r l) i0) (pos_Rl (cons r l) (S i0))) (pos_Rl x i0)
i:nat
H3:(i < Init.Nat.pred (Rlength (cons r l)))%nat
H5:forall x1 : R, pos_Rl (cons r l) i < x1 < pos_Rl (cons r l) (S i) -> f x1 = pos_Rl x i
x0:R
H6:pos_Rl (cons r l) i < x0 < pos_Rl (cons r l) (S i)
Hrecl:ordered_Rlist l -> pos_Rl l 0 = Rmin a b -> pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b -> Rlength l = S (Rlength x) -> (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i0) (pos_Rl l (S i0))) (pos_Rl x i0)) -> (i < Init.Nat.pred (Rlength l))%nat -> (forall x1 : R, pos_Rl l i < x1 < pos_Rl l (S i) -> f x1 = pos_Rl x i) -> pos_Rl l i < x0 < pos_Rl l (S i) -> f x0 = pos_Rl (FF l f) i
(i < Rlength (mid_Rlist l r))%nat
a, b:R
f:R -> R
r:R
l, x:Rlist
H:ordered_Rlist (cons r l)
H1:pos_Rl (cons r l) 0 = Rmin a b
H0:pos_Rl (cons r l) (Init.Nat.pred (Rlength (cons r l))) = Rmax a b
H2:Rlength (cons r l) = S (Rlength x)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r l) i0) (pos_Rl (cons r l) (S i0))) (pos_Rl x i0)
i:nat
H3:(i < Init.Nat.pred (Rlength (cons r l)))%nat
H5:forall x1 : R, pos_Rl (cons r l) i < x1 < pos_Rl (cons r l) (S i) -> f x1 = pos_Rl x i
x0:R
H6:pos_Rl (cons r l) i < x0 < pos_Rl (cons r l) (S i)
Hrecl:ordered_Rlist l -> pos_Rl l 0 = Rmin a b -> pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b -> Rlength l = S (Rlength x) -> (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i0) (pos_Rl l (S i0))) (pos_Rl x i0)) -> (i < Init.Nat.pred (Rlength l))%nat -> (forall x1 : R, pos_Rl l i < x1 < pos_Rl l (S i) -> f x1 = pos_Rl x i) -> pos_Rl l i < x0 < pos_Rl l (S i) -> f x0 = pos_Rl (FF l f) i

pos_Rl x i = pos_Rl x i
a, b:R
f:R -> R
r:R
l, x:Rlist
H:ordered_Rlist (cons r l)
H1:pos_Rl (cons r l) 0 = Rmin a b
H0:pos_Rl (cons r l) (Init.Nat.pred (Rlength (cons r l))) = Rmax a b
H2:Rlength (cons r l) = S (Rlength x)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r l) i0) (pos_Rl (cons r l) (S i0))) (pos_Rl x i0)
i:nat
H3:(i < Init.Nat.pred (Rlength (cons r l)))%nat
H5:forall x1 : R, pos_Rl (cons r l) i < x1 < pos_Rl (cons r l) (S i) -> f x1 = pos_Rl x i
x0:R
H6:pos_Rl (cons r l) i < x0 < pos_Rl (cons r l) (S i)
Hrecl:ordered_Rlist l -> pos_Rl l 0 = Rmin a b -> pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b -> Rlength l = S (Rlength x) -> (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i0) (pos_Rl l (S i0))) (pos_Rl x i0)) -> (i < Init.Nat.pred (Rlength l))%nat -> (forall x1 : R, pos_Rl l i < x1 < pos_Rl l (S i) -> f x1 = pos_Rl x i) -> pos_Rl l i < x0 < pos_Rl l (S i) -> f x0 = pos_Rl (FF l f) i
pos_Rl (cons r l) i < (pos_Rl (cons r l) i + pos_Rl (cons r l) (S i)) / 2 < pos_Rl (cons r l) (S i)
a, b:R
f:R -> R
r:R
l, x:Rlist
H:ordered_Rlist (cons r l)
H1:pos_Rl (cons r l) 0 = Rmin a b
H0:pos_Rl (cons r l) (Init.Nat.pred (Rlength (cons r l))) = Rmax a b
H2:Rlength (cons r l) = S (Rlength x)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r l) i0) (pos_Rl (cons r l) (S i0))) (pos_Rl x i0)
i:nat
H3:(i < Init.Nat.pred (Rlength (cons r l)))%nat
H5:forall x1 : R, pos_Rl (cons r l) i < x1 < pos_Rl (cons r l) (S i) -> f x1 = pos_Rl x i
x0:R
H6:pos_Rl (cons r l) i < x0 < pos_Rl (cons r l) (S i)
Hrecl:ordered_Rlist l -> pos_Rl l 0 = Rmin a b -> pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b -> Rlength l = S (Rlength x) -> (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i0) (pos_Rl l (S i0))) (pos_Rl x i0)) -> (i < Init.Nat.pred (Rlength l))%nat -> (forall x1 : R, pos_Rl l i < x1 < pos_Rl l (S i) -> f x1 = pos_Rl x i) -> pos_Rl l i < x0 < pos_Rl l (S i) -> f x0 = pos_Rl (FF l f) i
(i < Rlength (mid_Rlist l r))%nat
a, b:R
f:R -> R
r:R
l, x:Rlist
H:ordered_Rlist (cons r l)
H1:pos_Rl (cons r l) 0 = Rmin a b
H0:pos_Rl (cons r l) (Init.Nat.pred (Rlength (cons r l))) = Rmax a b
H2:Rlength (cons r l) = S (Rlength x)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r l) i0) (pos_Rl (cons r l) (S i0))) (pos_Rl x i0)
i:nat
H3:(i < Init.Nat.pred (Rlength (cons r l)))%nat
H5:forall x1 : R, pos_Rl (cons r l) i < x1 < pos_Rl (cons r l) (S i) -> f x1 = pos_Rl x i
x0:R
H6:pos_Rl (cons r l) i < x0 < pos_Rl (cons r l) (S i)
Hrecl:ordered_Rlist l -> pos_Rl l 0 = Rmin a b -> pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b -> Rlength l = S (Rlength x) -> (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i0) (pos_Rl l (S i0))) (pos_Rl x i0)) -> (i < Init.Nat.pred (Rlength l))%nat -> (forall x1 : R, pos_Rl l i < x1 < pos_Rl l (S i) -> f x1 = pos_Rl x i) -> pos_Rl l i < x0 < pos_Rl l (S i) -> f x0 = pos_Rl (FF l f) i

pos_Rl (cons r l) i < (pos_Rl (cons r l) i + pos_Rl (cons r l) (S i)) / 2 < pos_Rl (cons r l) (S i)
a, b:R
f:R -> R
r:R
l, x:Rlist
H:ordered_Rlist (cons r l)
H1:pos_Rl (cons r l) 0 = Rmin a b
H0:pos_Rl (cons r l) (Init.Nat.pred (Rlength (cons r l))) = Rmax a b
H2:Rlength (cons r l) = S (Rlength x)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r l) i0) (pos_Rl (cons r l) (S i0))) (pos_Rl x i0)
i:nat
H3:(i < Init.Nat.pred (Rlength (cons r l)))%nat
H5:forall x1 : R, pos_Rl (cons r l) i < x1 < pos_Rl (cons r l) (S i) -> f x1 = pos_Rl x i
x0:R
H6:pos_Rl (cons r l) i < x0 < pos_Rl (cons r l) (S i)
Hrecl:ordered_Rlist l -> pos_Rl l 0 = Rmin a b -> pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b -> Rlength l = S (Rlength x) -> (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i0) (pos_Rl l (S i0))) (pos_Rl x i0)) -> (i < Init.Nat.pred (Rlength l))%nat -> (forall x1 : R, pos_Rl l i < x1 < pos_Rl l (S i) -> f x1 = pos_Rl x i) -> pos_Rl l i < x0 < pos_Rl l (S i) -> f x0 = pos_Rl (FF l f) i
(i < Rlength (mid_Rlist l r))%nat
a, b:R
f:R -> R
r:R
l, x:Rlist
H:ordered_Rlist (cons r l)
H1:pos_Rl (cons r l) 0 = Rmin a b
H0:pos_Rl (cons r l) (Init.Nat.pred (Rlength (cons r l))) = Rmax a b
H2:Rlength (cons r l) = S (Rlength x)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r l) i0) (pos_Rl (cons r l) (S i0))) (pos_Rl x i0)
i:nat
H3:(i < Init.Nat.pred (Rlength (cons r l)))%nat
H5:forall x1 : R, pos_Rl (cons r l) i < x1 < pos_Rl (cons r l) (S i) -> f x1 = pos_Rl x i
x0:R
H6:pos_Rl (cons r l) i < x0 < pos_Rl (cons r l) (S i)
Hrecl:ordered_Rlist l -> pos_Rl l 0 = Rmin a b -> pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b -> Rlength l = S (Rlength x) -> (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i0) (pos_Rl l (S i0))) (pos_Rl x i0)) -> (i < Init.Nat.pred (Rlength l))%nat -> (forall x1 : R, pos_Rl l i < x1 < pos_Rl l (S i) -> f x1 = pos_Rl x i) -> pos_Rl l i < x0 < pos_Rl l (S i) -> f x0 = pos_Rl (FF l f) i

pos_Rl (cons r l) i < (pos_Rl (cons r l) i + pos_Rl (cons r l) (S i)) / 2
a, b:R
f:R -> R
r:R
l, x:Rlist
H:ordered_Rlist (cons r l)
H1:pos_Rl (cons r l) 0 = Rmin a b
H0:pos_Rl (cons r l) (Init.Nat.pred (Rlength (cons r l))) = Rmax a b
H2:Rlength (cons r l) = S (Rlength x)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r l) i0) (pos_Rl (cons r l) (S i0))) (pos_Rl x i0)
i:nat
H3:(i < Init.Nat.pred (Rlength (cons r l)))%nat
H5:forall x1 : R, pos_Rl (cons r l) i < x1 < pos_Rl (cons r l) (S i) -> f x1 = pos_Rl x i
x0:R
H6:pos_Rl (cons r l) i < x0 < pos_Rl (cons r l) (S i)
Hrecl:ordered_Rlist l -> pos_Rl l 0 = Rmin a b -> pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b -> Rlength l = S (Rlength x) -> (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i0) (pos_Rl l (S i0))) (pos_Rl x i0)) -> (i < Init.Nat.pred (Rlength l))%nat -> (forall x1 : R, pos_Rl l i < x1 < pos_Rl l (S i) -> f x1 = pos_Rl x i) -> pos_Rl l i < x0 < pos_Rl l (S i) -> f x0 = pos_Rl (FF l f) i
(pos_Rl (cons r l) i + pos_Rl (cons r l) (S i)) / 2 < pos_Rl (cons r l) (S i)
a, b:R
f:R -> R
r:R
l, x:Rlist
H:ordered_Rlist (cons r l)
H1:pos_Rl (cons r l) 0 = Rmin a b
H0:pos_Rl (cons r l) (Init.Nat.pred (Rlength (cons r l))) = Rmax a b
H2:Rlength (cons r l) = S (Rlength x)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r l) i0) (pos_Rl (cons r l) (S i0))) (pos_Rl x i0)
i:nat
H3:(i < Init.Nat.pred (Rlength (cons r l)))%nat
H5:forall x1 : R, pos_Rl (cons r l) i < x1 < pos_Rl (cons r l) (S i) -> f x1 = pos_Rl x i
x0:R
H6:pos_Rl (cons r l) i < x0 < pos_Rl (cons r l) (S i)
Hrecl:ordered_Rlist l -> pos_Rl l 0 = Rmin a b -> pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b -> Rlength l = S (Rlength x) -> (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i0) (pos_Rl l (S i0))) (pos_Rl x i0)) -> (i < Init.Nat.pred (Rlength l))%nat -> (forall x1 : R, pos_Rl l i < x1 < pos_Rl l (S i) -> f x1 = pos_Rl x i) -> pos_Rl l i < x0 < pos_Rl l (S i) -> f x0 = pos_Rl (FF l f) i
(i < Rlength (mid_Rlist l r))%nat
a, b:R
f:R -> R
r:R
l, x:Rlist
H:ordered_Rlist (cons r l)
H1:pos_Rl (cons r l) 0 = Rmin a b
H0:pos_Rl (cons r l) (Init.Nat.pred (Rlength (cons r l))) = Rmax a b
H2:Rlength (cons r l) = S (Rlength x)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r l) i0) (pos_Rl (cons r l) (S i0))) (pos_Rl x i0)
i:nat
H3:(i < Init.Nat.pred (Rlength (cons r l)))%nat
H5:forall x1 : R, pos_Rl (cons r l) i < x1 < pos_Rl (cons r l) (S i) -> f x1 = pos_Rl x i
x0:R
H6:pos_Rl (cons r l) i < x0 < pos_Rl (cons r l) (S i)
Hrecl:ordered_Rlist l -> pos_Rl l 0 = Rmin a b -> pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b -> Rlength l = S (Rlength x) -> (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i0) (pos_Rl l (S i0))) (pos_Rl x i0)) -> (i < Init.Nat.pred (Rlength l))%nat -> (forall x1 : R, pos_Rl l i < x1 < pos_Rl l (S i) -> f x1 = pos_Rl x i) -> pos_Rl l i < x0 < pos_Rl l (S i) -> f x0 = pos_Rl (FF l f) i

(pos_Rl (cons r l) i + pos_Rl (cons r l) (S i)) / 2 < pos_Rl (cons r l) (S i)
a, b:R
f:R -> R
r:R
l, x:Rlist
H:ordered_Rlist (cons r l)
H1:pos_Rl (cons r l) 0 = Rmin a b
H0:pos_Rl (cons r l) (Init.Nat.pred (Rlength (cons r l))) = Rmax a b
H2:Rlength (cons r l) = S (Rlength x)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r l) i0) (pos_Rl (cons r l) (S i0))) (pos_Rl x i0)
i:nat
H3:(i < Init.Nat.pred (Rlength (cons r l)))%nat
H5:forall x1 : R, pos_Rl (cons r l) i < x1 < pos_Rl (cons r l) (S i) -> f x1 = pos_Rl x i
x0:R
H6:pos_Rl (cons r l) i < x0 < pos_Rl (cons r l) (S i)
Hrecl:ordered_Rlist l -> pos_Rl l 0 = Rmin a b -> pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b -> Rlength l = S (Rlength x) -> (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i0) (pos_Rl l (S i0))) (pos_Rl x i0)) -> (i < Init.Nat.pred (Rlength l))%nat -> (forall x1 : R, pos_Rl l i < x1 < pos_Rl l (S i) -> f x1 = pos_Rl x i) -> pos_Rl l i < x0 < pos_Rl l (S i) -> f x0 = pos_Rl (FF l f) i
(i < Rlength (mid_Rlist l r))%nat
a, b:R
f:R -> R
r:R
l, x:Rlist
H:ordered_Rlist (cons r l)
H1:pos_Rl (cons r l) 0 = Rmin a b
H0:pos_Rl (cons r l) (Init.Nat.pred (Rlength (cons r l))) = Rmax a b
H2:Rlength (cons r l) = S (Rlength x)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r l)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r l) i0) (pos_Rl (cons r l) (S i0))) (pos_Rl x i0)
i:nat
H3:(i < Init.Nat.pred (Rlength (cons r l)))%nat
H5:forall x1 : R, pos_Rl (cons r l) i < x1 < pos_Rl (cons r l) (S i) -> f x1 = pos_Rl x i
x0:R
H6:pos_Rl (cons r l) i < x0 < pos_Rl (cons r l) (S i)
Hrecl:ordered_Rlist l -> pos_Rl l 0 = Rmin a b -> pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b -> Rlength l = S (Rlength x) -> (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i0) (pos_Rl l (S i0))) (pos_Rl x i0)) -> (i < Init.Nat.pred (Rlength l))%nat -> (forall x1 : R, pos_Rl l i < x1 < pos_Rl l (S i) -> f x1 = pos_Rl x i) -> pos_Rl l i < x0 < pos_Rl l (S i) -> f x0 = pos_Rl (FF l f) i

(i < Rlength (mid_Rlist l r))%nat
rewrite RList_P14; simpl in H3; apply H3. Qed.

forall (a b : R) (f g : R -> R) (lf lg : Rlist), a <= b -> is_subdivision f a b lf -> is_subdivision g a b lg -> is_subdivision f a b (cons_ORlist lf lg)

forall (a b : R) (f g : R -> R) (lf lg : Rlist), a <= b -> is_subdivision f a b lf -> is_subdivision g a b lg -> is_subdivision f a b (cons_ORlist lf lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0:Rlist
p:adapted_couple g a b lg lg0
lf0:Rlist
p0:adapted_couple f a b lf lf0

Rmin a b = a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0:Rlist
p:adapted_couple g a b lg lg0
lf0:Rlist
p0:adapted_couple f a b lf lf0
Hyp_min:Rmin a b = a
{l0 : Rlist & adapted_couple f a b (cons_ORlist lf lg) l0}
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0:Rlist
p:adapted_couple g a b lg lg0
lf0:Rlist
p0:adapted_couple f a b lf lf0
Hyp_min:Rmin a b = a

{l0 : Rlist & adapted_couple f a b (cons_ORlist lf lg) l0}
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0:Rlist
p:adapted_couple g a b lg lg0
lf0:Rlist
p0:adapted_couple f a b lf lf0
Hyp_min:Rmin a b = a

Rmax a b = b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0:Rlist
p:adapted_couple g a b lg lg0
lf0:Rlist
p0:adapted_couple f a b lf lf0
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
{l0 : Rlist & adapted_couple f a b (cons_ORlist lf lg) l0}
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0:Rlist
p:adapted_couple g a b lg lg0
lf0:Rlist
p0:adapted_couple f a b lf lf0
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b

{l0 : Rlist & adapted_couple f a b (cons_ORlist lf lg) l0}
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)

ordered_Rlist (cons_ORlist lf lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) 0 = Rmin a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)

pos_Rl (cons_ORlist lf lg) 0 = Rmin a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)

a <= pos_Rl (cons_ORlist lf lg) 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist nil
H6:pos_Rl nil 0 = a
H5:pos_Rl nil (Init.Nat.pred (Rlength nil)) = b
H7:Rlength nil = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength nil))%nat -> constant_D_eq f (open_interval (pos_Rl nil i) (pos_Rl nil (S i))) (pos_Rl lf0 i)

a <= pos_Rl (cons_ORlist nil lg) 0
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> a <= pos_Rl (cons_ORlist lf lg) 0
a <= pos_Rl (cons_ORlist (cons r lf) lg) 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> a <= pos_Rl (cons_ORlist lf lg) 0

a <= pos_Rl (cons_ORlist (cons r lf) lg) 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> a <= pos_Rl (cons_ORlist lf lg) 0

In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons_ORlist (cons r lf) lg)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> a <= pos_Rl (cons_ORlist lf lg) 0
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons_ORlist (cons r lf) lg)
a <= pos_Rl (cons_ORlist (cons r lf) lg) 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> a <= pos_Rl (cons_ORlist lf lg) 0
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons_ORlist (cons r lf) lg)

a <= pos_Rl (cons_ORlist (cons r lf) lg) 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> a <= pos_Rl (cons_ORlist lf lg) 0
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons_ORlist (cons r lf) lg)
H12:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons_ORlist (cons r lf) lg) -> In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) 0) lg
H13:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) 0) lg
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons r lf)

a <= pos_Rl (cons_ORlist (cons r lf) lg) 0
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> a <= pos_Rl (cons_ORlist lf lg) 0
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons_ORlist (cons r lf) lg)
H12:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons_ORlist (cons r lf) lg) -> In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) 0) lg
H13:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) 0) lg
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) lg
a <= pos_Rl (cons_ORlist (cons r lf) lg) 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> a <= pos_Rl (cons_ORlist lf lg) 0
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons_ORlist (cons r lf) lg)
H12:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons_ORlist (cons r lf) lg) -> In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) 0) lg
H13:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) 0) lg
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) lg

a <= pos_Rl (cons_ORlist (cons r lf) lg) 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)

pos_Rl (cons_ORlist lf lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist nil
H6:pos_Rl nil 0 = a
H5:pos_Rl nil (Init.Nat.pred (Rlength nil)) = b
H7:Rlength nil = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength nil))%nat -> constant_D_eq f (open_interval (pos_Rl nil i) (pos_Rl nil (S i))) (pos_Rl lf0 i)

pos_Rl (cons_ORlist nil lg) 0 <= a
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) 0 <= a
pos_Rl (cons_ORlist (cons r lf) lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) 0 <= a

pos_Rl (cons_ORlist (cons r lf) lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) 0 <= a

In a (cons_ORlist (cons r lf) lg)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) 0 <= a
H8:In a (cons_ORlist (cons r lf) lg)
pos_Rl (cons_ORlist (cons r lf) lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) 0 <= a
H8:In a (cons_ORlist (cons r lf) lg)

pos_Rl (cons_ORlist (cons r lf) lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)

pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)

pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist nil
H6:pos_Rl nil 0 = a
H5:pos_Rl nil (Init.Nat.pred (Rlength nil)) = b
H7:Rlength nil = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength nil))%nat -> constant_D_eq f (open_interval (pos_Rl nil i) (pos_Rl nil (S i))) (pos_Rl lf0 i)

pos_Rl (cons_ORlist nil lg) (Init.Nat.pred (Rlength (cons_ORlist nil lg))) <= b
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b

pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b

In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg)
pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg)

pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg)
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg) -> In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg

pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg)
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg) -> In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H11:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H12:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf)

pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) <= b
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg)
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg) -> In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H11:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H12:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg)
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg) -> In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H11:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H12:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg

pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg)
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg) -> In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H11:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H12:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H13:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg -> exists i : nat, pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) = pos_Rl lg i /\ (i < Rlength lg)%nat
H14:exists i : nat, pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) = pos_Rl lg i /\ (i < Rlength lg)%nat
x:nat
H15:pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) = pos_Rl lg x
H16:(x < Rlength lg)%nat

pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg)
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg) -> In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H11:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H12:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H13:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg -> exists i : nat, pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) = pos_Rl lg i /\ (i < Rlength lg)%nat
H14:exists i : nat, pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) = pos_Rl lg i /\ (i < Rlength lg)%nat
x:nat
H15:pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) = pos_Rl lg x
H16:(x < Rlength lg)%nat

Rlength lg = S (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg)
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg) -> In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H11:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H12:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H13:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg -> exists i : nat, pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) = pos_Rl lg i /\ (i < Rlength lg)%nat
H14:exists i : nat, pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) = pos_Rl lg i /\ (i < Rlength lg)%nat
x:nat
H15:pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) = pos_Rl lg x
H16:(x < Rlength lg)%nat
H17:Rlength lg = S (Init.Nat.pred (Rlength lg))
pos_Rl lg x <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg)
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg) -> In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H11:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H12:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H13:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg -> exists i : nat, pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) = pos_Rl lg i /\ (i < Rlength lg)%nat
H14:exists i : nat, pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) = pos_Rl lg i /\ (i < Rlength lg)%nat
x:nat
H15:pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) = pos_Rl lg x
H16:(x < Rlength lg)%nat
H17:Rlength lg = S (Init.Nat.pred (Rlength lg))

pos_Rl lg x <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)

b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist nil
H6:pos_Rl nil 0 = a
H5:pos_Rl nil (Init.Nat.pred (Rlength nil)) = b
H7:Rlength nil = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength nil))%nat -> constant_D_eq f (open_interval (pos_Rl nil i) (pos_Rl nil (S i))) (pos_Rl lf0 i)

b <= pos_Rl (cons_ORlist nil lg) (Init.Nat.pred (Rlength (cons_ORlist nil lg)))
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
b <= pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))

b <= pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))

In b (cons_ORlist (cons r lf) lg)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
H8:In b (cons_ORlist (cons r lf) lg)
b <= pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
H8:In b (cons_ORlist (cons r lf) lg)

b <= pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)

Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) f))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)

forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) f) i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

(exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l) -> f x = pos_Rl (FF (cons_ORlist lf lg) f) i
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0

exists (r : R) (r0 : Rlist), cons_ORlist lf lg = cons r r0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
Hyp_cons:exists (r : R) (r0 : Rlist), cons_ORlist lf lg = cons r r0
f x = pos_Rl (FF (cons_ORlist lf lg) f) i
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
Hyp_cons:exists (r : R) (r0 : Rlist), cons_ORlist lf lg = cons r r0

f x = pos_Rl (FF (cons_ORlist lf lg) f) i
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0

f x = f (pos_Rl (mid_Rlist r0 r) i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Rlength (mid_Rlist r0 r))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0

f x = f ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Rlength (mid_Rlist r0 r))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)

f x = f ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i = pos_Rl (cons_ORlist lf lg) (S i)
f x = f ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Rlength (mid_Rlist r0 r))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:forall x1 : R, pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i) -> f x1 = x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)

pos_Rl (cons_ORlist lf lg) i < (pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2 < pos_Rl (cons_ORlist lf lg) (S i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:forall x1 : R, pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i) -> f x1 = x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)
H15:pos_Rl (cons_ORlist lf lg) i < (pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2 < pos_Rl (cons_ORlist lf lg) (S i)
x0 = f ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i = pos_Rl (cons_ORlist lf lg) (S i)
f x = f ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Rlength (mid_Rlist r0 r))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:forall x1 : R, pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i) -> f x1 = x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)

pos_Rl (cons_ORlist lf lg) i < (pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:forall x1 : R, pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i) -> f x1 = x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)
(pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2 < pos_Rl (cons_ORlist lf lg) (S i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:forall x1 : R, pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i) -> f x1 = x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)
H15:pos_Rl (cons_ORlist lf lg) i < (pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2 < pos_Rl (cons_ORlist lf lg) (S i)
x0 = f ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i = pos_Rl (cons_ORlist lf lg) (S i)
f x = f ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Rlength (mid_Rlist r0 r))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:forall x1 : R, pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i) -> f x1 = x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)

(pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2 < pos_Rl (cons_ORlist lf lg) (S i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:forall x1 : R, pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i) -> f x1 = x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)
H15:pos_Rl (cons_ORlist lf lg) i < (pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2 < pos_Rl (cons_ORlist lf lg) (S i)
x0 = f ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i = pos_Rl (cons_ORlist lf lg) (S i)
f x = f ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Rlength (mid_Rlist r0 r))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:forall x1 : R, pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i) -> f x1 = x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)
H15:pos_Rl (cons_ORlist lf lg) i < (pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2 < pos_Rl (cons_ORlist lf lg) (S i)

x0 = f ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i = pos_Rl (cons_ORlist lf lg) (S i)
f x = f ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Rlength (mid_Rlist r0 r))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i = pos_Rl (cons_ORlist lf lg) (S i)

f x = f ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Rlength (mid_Rlist r0 r))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0

(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Rlength (mid_Rlist r0 r))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0

(i < Rlength (mid_Rlist r0 r))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

a < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

a <= pos_Rl (cons_ORlist lf lg) i
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl (cons_ORlist lf lg) i < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

pos_Rl (cons_ORlist lf lg) 0 <= pos_Rl (cons_ORlist lf lg) i
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
ordered_Rlist lf
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
ordered_Rlist lg
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl lf 0 = pos_Rl lg 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl (cons_ORlist lf lg) i < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

ordered_Rlist (cons_ORlist lf lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(0 <= i)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(i < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
ordered_Rlist lf
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
ordered_Rlist lg
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl lf 0 = pos_Rl lg 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl (cons_ORlist lf lg) i < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

(0 <= i)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(i < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
ordered_Rlist lf
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
ordered_Rlist lg
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl lf 0 = pos_Rl lg 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl (cons_ORlist lf lg) i < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

(i < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
ordered_Rlist lf
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
ordered_Rlist lg
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl lf 0 = pos_Rl lg 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl (cons_ORlist lf lg) i < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

ordered_Rlist lf
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
ordered_Rlist lg
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl lf 0 = pos_Rl lg 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl (cons_ORlist lf lg) i < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

ordered_Rlist lg
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl lf 0 = pos_Rl lg 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl (cons_ORlist lf lg) i < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

pos_Rl lf 0 = pos_Rl lg 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl (cons_ORlist lf lg) i < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

pos_Rl (cons_ORlist lf lg) i < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl (cons_ORlist lf lg) (S i) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

pos_Rl (cons_ORlist lf lg) (S i) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl lf (Init.Nat.pred (Rlength lf)) = pos_Rl lg (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

ordered_Rlist (cons_ORlist lf lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(S i <= Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(Init.Nat.pred (Rlength (cons_ORlist lf lg)) < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl lf (Init.Nat.pred (Rlength lf)) = pos_Rl lg (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

(S i <= Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(Init.Nat.pred (Rlength (cons_ORlist lf lg)) < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl lf (Init.Nat.pred (Rlength lf)) = pos_Rl lg (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

(Init.Nat.pred (Rlength (cons_ORlist lf lg)) < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl lf (Init.Nat.pred (Rlength lf)) = pos_Rl lg (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

pos_Rl lf (Init.Nat.pred (Rlength lf)) = pos_Rl lg (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b

exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop

Nbound I
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I

exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I

exists n : nat, I n
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I

pos_Rl lf 0 <= pos_Rl (cons_ORlist lf lg) i
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
(0 < Rlength lf)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I

pos_Rl lf 0 <= pos_Rl (cons_ORlist lf lg) 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
pos_Rl (cons_ORlist lf lg) 0 <= pos_Rl (cons_ORlist lf lg) i
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
(0 < Rlength lf)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I

pos_Rl (cons_ORlist lf lg) 0 = pos_Rl lf 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
pos_Rl (cons_ORlist lf lg) 0 <= pos_Rl (cons_ORlist lf lg) i
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
(0 < Rlength lf)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I

pos_Rl (cons_ORlist lf lg) 0 <= pos_Rl (cons_ORlist lf lg) i
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
(0 < Rlength lf)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H14:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

ordered_Rlist (cons_ORlist lf lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H14:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(0 <= i)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H14:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(i < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
(0 < Rlength lf)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H14:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

(0 <= i)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H14:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(i < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
(0 < Rlength lf)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H14:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

(i < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
(0 < Rlength lf)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H14:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H14:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(Init.Nat.pred (Rlength (cons_ORlist lf lg)) < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
(0 < Rlength lf)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H14:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

(Init.Nat.pred (Rlength (cons_ORlist lf lg)) < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
(0 < Rlength lf)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I

(0 < Rlength lf)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n

exists l : R, constant_D_eq f (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)

(x0 < Init.Nat.pred (Rlength lf))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
f x1 = pos_Rl lf0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat

(S x0 < Rlength lf)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
Rlength lf = S (Init.Nat.pred (Rlength lf))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
f x1 = pos_Rl lf0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf

(S x0 < S x0)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
m:nat
H20:(S x0 <= m)%nat
H19:S m = Rlength lf
(S x0 < S m)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
Rlength lf = S (Init.Nat.pred (Rlength lf))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
f x1 = pos_Rl lf0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf

(S x0 < S x0)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
Rlength lf = S (Init.Nat.pred (Rlength lf))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
f x1 = pos_Rl lf0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf

x0 = Init.Nat.pred (Rlength lf) -> (S x0 < S x0)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
x0 = Init.Nat.pred (Rlength lf)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
Rlength lf = S (Init.Nat.pred (Rlength lf))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
f x1 = pos_Rl lf0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
H19:x0 = Init.Nat.pred (Rlength lf)

pos_Rl (cons_ORlist lf lg) i < b -> (S x0 < S x0)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
H19:x0 = Init.Nat.pred (Rlength lf)
pos_Rl (cons_ORlist lf lg) i < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
x0 = Init.Nat.pred (Rlength lf)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
Rlength lf = S (Init.Nat.pred (Rlength lf))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
f x1 = pos_Rl lf0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
H19:x0 = Init.Nat.pred (Rlength lf)

pos_Rl (cons_ORlist lf lg) i < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
x0 = Init.Nat.pred (Rlength lf)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
Rlength lf = S (Init.Nat.pred (Rlength lf))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
f x1 = pos_Rl lf0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
H19:x0 = Init.Nat.pred (Rlength lf)

pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
H19:x0 = Init.Nat.pred (Rlength lf)
pos_Rl (cons_ORlist lf lg) (S i) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
x0 = Init.Nat.pred (Rlength lf)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
Rlength lf = S (Init.Nat.pred (Rlength lf))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
f x1 = pos_Rl lf0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
H19:x0 = Init.Nat.pred (Rlength lf)

pos_Rl (cons_ORlist lf lg) (S i) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
x0 = Init.Nat.pred (Rlength lf)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
Rlength lf = S (Init.Nat.pred (Rlength lf))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
f x1 = pos_Rl lf0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
H19:x0 = Init.Nat.pred (Rlength lf)

pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
H19:x0 = Init.Nat.pred (Rlength lf)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= pos_Rl lf (Init.Nat.pred (Rlength lf))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
x0 = Init.Nat.pred (Rlength lf)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
Rlength lf = S (Init.Nat.pred (Rlength lf))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
f x1 = pos_Rl lf0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
H19:x0 = Init.Nat.pred (Rlength lf)
H21:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H22:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

ordered_Rlist (cons_ORlist lf lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
H19:x0 = Init.Nat.pred (Rlength lf)
H21:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H22:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(S i <= Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
H19:x0 = Init.Nat.pred (Rlength lf)
H21:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H22:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(Init.Nat.pred (Rlength (cons_ORlist lf lg)) < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
H19:x0 = Init.Nat.pred (Rlength lf)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= pos_Rl lf (Init.Nat.pred (Rlength lf))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
x0 = Init.Nat.pred (Rlength lf)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
Rlength lf = S (Init.Nat.pred (Rlength lf))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
f x1 = pos_Rl lf0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
H19:x0 = Init.Nat.pred (Rlength lf)
H21:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H22:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

(S i <= Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
H19:x0 = Init.Nat.pred (Rlength lf)
H21:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H22:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(Init.Nat.pred (Rlength (cons_ORlist lf lg)) < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
H19:x0 = Init.Nat.pred (Rlength lf)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= pos_Rl lf (Init.Nat.pred (Rlength lf))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
x0 = Init.Nat.pred (Rlength lf)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
Rlength lf = S (Init.Nat.pred (Rlength lf))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
f x1 = pos_Rl lf0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
H19:x0 = Init.Nat.pred (Rlength lf)
H21:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H22:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

(Init.Nat.pred (Rlength (cons_ORlist lf lg)) < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
H19:x0 = Init.Nat.pred (Rlength lf)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= pos_Rl lf (Init.Nat.pred (Rlength lf))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
x0 = Init.Nat.pred (Rlength lf)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
Rlength lf = S (Init.Nat.pred (Rlength lf))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
f x1 = pos_Rl lf0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
H19:x0 = Init.Nat.pred (Rlength lf)

pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= pos_Rl lf (Init.Nat.pred (Rlength lf))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf
x0 = Init.Nat.pred (Rlength lf)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
Rlength lf = S (Init.Nat.pred (Rlength lf))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
f x1 = pos_Rl lf0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
H20:S x0 = Rlength lf

x0 = Init.Nat.pred (Rlength lf)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat
Rlength lf = S (Init.Nat.pred (Rlength lf))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
f x1 = pos_Rl lf0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lf)%nat

Rlength lf = S (Init.Nat.pred (Rlength lf))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
f x1 = pos_Rl lf0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat

f x1 = pos_Rl lf0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0

pos_Rl lf0 x0 = pos_Rl lf0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
pos_Rl lf x0 < x1 < pos_Rl lf (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0

pos_Rl lf x0 < x1 < pos_Rl lf (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lf)%nat

pos_Rl lf x0 < x1
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lf)%nat
x1 < pos_Rl lf (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lf)%nat

x1 < pos_Rl lf (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lf)%nat

pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lf (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lf)%nat

(S x0 < Rlength lf)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lf)%nat
H22:(S x0 < Rlength lf)%nat
pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lf (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lf)%nat
H22:(S x0 < Rlength lf)%nat

pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lf (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lf)%nat
H22:(S x0 < Rlength lf)%nat
a0:pos_Rl lf (S x0) <= pos_Rl (cons_ORlist lf lg) i

pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lf (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lf)%nat
H22:(S x0 < Rlength lf)%nat
a0:~ pos_Rl lf (S x0) <= pos_Rl (cons_ORlist lf lg) i
pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lf (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lf)%nat
H22:(S x0 < Rlength lf)%nat
a0:pos_Rl lf (S x0) <= pos_Rl (cons_ORlist lf lg) i

(S x0 <= x0)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lf)%nat
H22:(S x0 < Rlength lf)%nat
a0:pos_Rl lf (S x0) <= pos_Rl (cons_ORlist lf lg) i
H23:(S x0 <= x0)%nat
pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lf (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lf)%nat
H22:(S x0 < Rlength lf)%nat
a0:~ pos_Rl lf (S x0) <= pos_Rl (cons_ORlist lf lg) i
pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lf (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lf)%nat
H22:(S x0 < Rlength lf)%nat
a0:pos_Rl lf (S x0) <= pos_Rl (cons_ORlist lf lg) i
H23:(S x0 <= x0)%nat

pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lf (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lf)%nat
H22:(S x0 < Rlength lf)%nat
a0:~ pos_Rl lf (S x0) <= pos_Rl (cons_ORlist lf lg) i
pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lf (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lf)%nat
H22:(S x0 < Rlength lf)%nat
a0:~ pos_Rl lf (S x0) <= pos_Rl (cons_ORlist lf lg) i

pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lf (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lf)%nat
H22:(S x0 < Rlength lf)%nat
a0:~ pos_Rl lf (S x0) <= pos_Rl (cons_ORlist lf lg) i

pos_Rl (cons_ORlist lf lg) i < pos_Rl lf (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lf)%nat
H22:(S x0 < Rlength lf)%nat
a0:~ pos_Rl lf (S x0) <= pos_Rl (cons_ORlist lf lg) i
H23:pos_Rl (cons_ORlist lf lg) i < pos_Rl lf (S x0)
pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lf (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lf)%nat
H22:(S x0 < Rlength lf)%nat
a0:~ pos_Rl lf (S x0) <= pos_Rl (cons_ORlist lf lg) i
H23:pos_Rl (cons_ORlist lf lg) i < pos_Rl lf (S x0)

pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lf (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lf)%nat
H22:(S x0 < Rlength lf)%nat
H23:pos_Rl (cons_ORlist lf lg) i < pos_Rl lf (S x0)

ordered_Rlist (cons_ORlist lf lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lf)%nat
H22:(S x0 < Rlength lf)%nat
H23:pos_Rl (cons_ORlist lf lg) i < pos_Rl lf (S x0)
In (pos_Rl lf (S x0)) (cons_ORlist lf lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lf j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lf)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf x0) (pos_Rl lf (S x0))) (pos_Rl lf0 x0)
H17:(x0 < Init.Nat.pred (Rlength lf))%nat
H18:forall x2 : R, pos_Rl lf x0 < x2 < pos_Rl lf (S x0) -> f x2 = pos_Rl lf0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lf x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lf)%nat
H22:(S x0 < Rlength lf)%nat
H23:pos_Rl (cons_ORlist lf lg) i < pos_Rl lf (S x0)

In (pos_Rl lf (S x0)) (cons_ORlist lf lg)
elim (RList_P9 lf lg (pos_Rl lf (S x0))); intros; apply H25; left; elim (RList_P3 lf (pos_Rl lf (S x0))); intros; apply H27; exists (S x0); split; [ reflexivity | apply H22 ]. Qed.

forall (a b : R) (f g : R -> R) (lf lg : Rlist), is_subdivision f a b lf -> is_subdivision g a b lg -> is_subdivision f a b (cons_ORlist lf lg)

forall (a b : R) (f g : R -> R) (lf lg : Rlist), is_subdivision f a b lf -> is_subdivision g a b lg -> is_subdivision f a b (cons_ORlist lf lg)
intros; case (Rle_dec a b); intro; [ apply StepFun_P22 with g; assumption | apply StepFun_P5; apply StepFun_P22 with g; [ auto with real | apply StepFun_P5; assumption | apply StepFun_P5; assumption ] ]. Qed.

forall (a b : R) (f g : R -> R) (lf lg : Rlist), a <= b -> is_subdivision f a b lf -> is_subdivision g a b lg -> is_subdivision g a b (cons_ORlist lf lg)

forall (a b : R) (f g : R -> R) (lf lg : Rlist), a <= b -> is_subdivision f a b lf -> is_subdivision g a b lg -> is_subdivision g a b (cons_ORlist lf lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0:Rlist
p:adapted_couple g a b lg lg0
lf0:Rlist
p0:adapted_couple f a b lf lf0

Rmin a b = a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0:Rlist
p:adapted_couple g a b lg lg0
lf0:Rlist
p0:adapted_couple f a b lf lf0
Hyp_min:Rmin a b = a
{l0 : Rlist & adapted_couple g a b (cons_ORlist lf lg) l0}
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0:Rlist
p:adapted_couple g a b lg lg0
lf0:Rlist
p0:adapted_couple f a b lf lf0
Hyp_min:Rmin a b = a

{l0 : Rlist & adapted_couple g a b (cons_ORlist lf lg) l0}
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0:Rlist
p:adapted_couple g a b lg lg0
lf0:Rlist
p0:adapted_couple f a b lf lf0
Hyp_min:Rmin a b = a

Rmax a b = b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0:Rlist
p:adapted_couple g a b lg lg0
lf0:Rlist
p0:adapted_couple f a b lf lf0
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
{l0 : Rlist & adapted_couple g a b (cons_ORlist lf lg) l0}
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0:Rlist
p:adapted_couple g a b lg lg0
lf0:Rlist
p0:adapted_couple f a b lf lf0
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b

{l0 : Rlist & adapted_couple g a b (cons_ORlist lf lg) l0}
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)

ordered_Rlist (cons_ORlist lf lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) 0 = Rmin a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)

pos_Rl (cons_ORlist lf lg) 0 = Rmin a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)

a <= pos_Rl (cons_ORlist lf lg) 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist nil
H6:pos_Rl nil 0 = a
H5:pos_Rl nil (Init.Nat.pred (Rlength nil)) = b
H7:Rlength nil = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength nil))%nat -> constant_D_eq f (open_interval (pos_Rl nil i) (pos_Rl nil (S i))) (pos_Rl lf0 i)

a <= pos_Rl (cons_ORlist nil lg) 0
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> a <= pos_Rl (cons_ORlist lf lg) 0
a <= pos_Rl (cons_ORlist (cons r lf) lg) 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> a <= pos_Rl (cons_ORlist lf lg) 0

a <= pos_Rl (cons_ORlist (cons r lf) lg) 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> a <= pos_Rl (cons_ORlist lf lg) 0

In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons_ORlist (cons r lf) lg)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> a <= pos_Rl (cons_ORlist lf lg) 0
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons_ORlist (cons r lf) lg)
a <= pos_Rl (cons_ORlist (cons r lf) lg) 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> a <= pos_Rl (cons_ORlist lf lg) 0
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons_ORlist (cons r lf) lg)

a <= pos_Rl (cons_ORlist (cons r lf) lg) 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> a <= pos_Rl (cons_ORlist lf lg) 0
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons_ORlist (cons r lf) lg)
H12:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons_ORlist (cons r lf) lg) -> In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) 0) lg
H13:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) 0) lg
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons r lf)

a <= pos_Rl (cons_ORlist (cons r lf) lg) 0
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> a <= pos_Rl (cons_ORlist lf lg) 0
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons_ORlist (cons r lf) lg)
H12:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons_ORlist (cons r lf) lg) -> In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) 0) lg
H13:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) 0) lg
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) lg
a <= pos_Rl (cons_ORlist (cons r lf) lg) 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> a <= pos_Rl (cons_ORlist lf lg) 0
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons_ORlist (cons r lf) lg)
H12:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons_ORlist (cons r lf) lg) -> In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) 0) lg
H13:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) 0) lg
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) 0) lg

a <= pos_Rl (cons_ORlist (cons r lf) lg) 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)

pos_Rl (cons_ORlist lf lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist nil
H6:pos_Rl nil 0 = a
H5:pos_Rl nil (Init.Nat.pred (Rlength nil)) = b
H7:Rlength nil = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength nil))%nat -> constant_D_eq f (open_interval (pos_Rl nil i) (pos_Rl nil (S i))) (pos_Rl lf0 i)

pos_Rl (cons_ORlist nil lg) 0 <= a
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) 0 <= a
pos_Rl (cons_ORlist (cons r lf) lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) 0 <= a

pos_Rl (cons_ORlist (cons r lf) lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) 0 <= a

In a (cons_ORlist (cons r lf) lg)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) 0 <= a
H8:In a (cons_ORlist (cons r lf) lg)
pos_Rl (cons_ORlist (cons r lf) lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) 0 <= a
H8:In a (cons_ORlist (cons r lf) lg)

pos_Rl (cons_ORlist (cons r lf) lg) 0 <= a
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)

pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) = Rmax a b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)

pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist nil
H6:pos_Rl nil 0 = a
H5:pos_Rl nil (Init.Nat.pred (Rlength nil)) = b
H7:Rlength nil = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength nil))%nat -> constant_D_eq f (open_interval (pos_Rl nil i) (pos_Rl nil (S i))) (pos_Rl lf0 i)

pos_Rl (cons_ORlist nil lg) (Init.Nat.pred (Rlength (cons_ORlist nil lg))) <= b
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b

pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b

In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg)
pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg)

pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg)
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg) -> In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H11:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H12:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf)

pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) <= b
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg)
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg) -> In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H11:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H12:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg)
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg) -> In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H11:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H12:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg

pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg)
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg) -> In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H11:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H12:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H13:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg -> exists i : nat, pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) = pos_Rl lg i /\ (i < Rlength lg)%nat
H14:exists i : nat, pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) = pos_Rl lg i /\ (i < Rlength lg)%nat
x:nat
H15:pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) = pos_Rl lg x
H16:(x < Rlength lg)%nat

Rlength lg = S (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg)
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg) -> In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H11:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H12:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H13:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg -> exists i : nat, pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) = pos_Rl lg i /\ (i < Rlength lg)%nat
H14:exists i : nat, pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) = pos_Rl lg i /\ (i < Rlength lg)%nat
x:nat
H15:pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) = pos_Rl lg x
H16:(x < Rlength lg)%nat
H17:Rlength lg = S (Init.Nat.pred (Rlength lg))
pos_Rl lg x <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= b
H8:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg)
H10:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons_ORlist (cons r lf) lg) -> In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H11:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) (cons r lf) \/ In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H12:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg
H13:In (pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))) lg -> exists i : nat, pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) = pos_Rl lg i /\ (i < Rlength lg)%nat
H14:exists i : nat, pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) = pos_Rl lg i /\ (i < Rlength lg)%nat
x:nat
H15:pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg))) = pos_Rl lg x
H16:(x < Rlength lg)%nat
H17:Rlength lg = S (Init.Nat.pred (Rlength lg))

pos_Rl lg x <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)

b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist nil
H6:pos_Rl nil 0 = a
H5:pos_Rl nil (Init.Nat.pred (Rlength nil)) = b
H7:Rlength nil = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength nil))%nat -> constant_D_eq f (open_interval (pos_Rl nil i) (pos_Rl nil (S i))) (pos_Rl lf0 i)

b <= pos_Rl (cons_ORlist nil lg) (Init.Nat.pred (Rlength (cons_ORlist nil lg)))
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
b <= pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))

b <= pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))

In b (cons_ORlist (cons r lf) lg)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
H8:In b (cons_ORlist (cons r lf) lg)
b <= pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
r:R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist (cons r lf)
H6:pos_Rl (cons r lf) 0 = a
H5:pos_Rl (cons r lf) (Init.Nat.pred (Rlength (cons r lf))) = b
H7:Rlength (cons r lf) = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength (cons r lf)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r lf) i) (pos_Rl (cons r lf) (S i))) (pos_Rl lf0 i)
Hreclf:ordered_Rlist lf -> pos_Rl lf 0 = a -> pos_Rl lf (Init.Nat.pred (Rlength lf)) = b -> Rlength lf = S (Rlength lf0) -> (forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)) -> b <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
H8:In b (cons_ORlist (cons r lf) lg)

b <= pos_Rl (cons_ORlist (cons r lf) lg) (Init.Nat.pred (Rlength (cons_ORlist (cons r lf) lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)

Rlength (cons_ORlist lf lg) = S (Rlength (FF (cons_ORlist lf lg) g))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg0 i)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i : nat, (i < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i) (pos_Rl lf (S i))) (pos_Rl lf0 i)

forall i : nat, (i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) (pos_Rl (FF (cons_ORlist lf lg) g) i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

(exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l) -> g x = pos_Rl (FF (cons_ORlist lf lg) g) i
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0

exists (r : R) (r0 : Rlist), cons_ORlist lf lg = cons r r0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
Hyp_cons:exists (r : R) (r0 : Rlist), cons_ORlist lf lg = cons r r0
g x = pos_Rl (FF (cons_ORlist lf lg) g) i
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
Hyp_cons:exists (r : R) (r0 : Rlist), cons_ORlist lf lg = cons r r0

g x = pos_Rl (FF (cons_ORlist lf lg) g) i
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0

g x = g (pos_Rl (mid_Rlist r0 r) i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Rlength (mid_Rlist r0 r))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0

g x = g ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Rlength (mid_Rlist r0 r))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)

g x = g ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i = pos_Rl (cons_ORlist lf lg) (S i)
g x = g ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Rlength (mid_Rlist r0 r))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:forall x1 : R, pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i) -> g x1 = x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)

pos_Rl (cons_ORlist lf lg) i < (pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2 < pos_Rl (cons_ORlist lf lg) (S i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:forall x1 : R, pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i) -> g x1 = x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)
H15:pos_Rl (cons_ORlist lf lg) i < (pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2 < pos_Rl (cons_ORlist lf lg) (S i)
x0 = g ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i = pos_Rl (cons_ORlist lf lg) (S i)
g x = g ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Rlength (mid_Rlist r0 r))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:forall x1 : R, pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i) -> g x1 = x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)

pos_Rl (cons_ORlist lf lg) i < (pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:forall x1 : R, pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i) -> g x1 = x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)
(pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2 < pos_Rl (cons_ORlist lf lg) (S i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:forall x1 : R, pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i) -> g x1 = x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)
H15:pos_Rl (cons_ORlist lf lg) i < (pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2 < pos_Rl (cons_ORlist lf lg) (S i)
x0 = g ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i = pos_Rl (cons_ORlist lf lg) (S i)
g x = g ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Rlength (mid_Rlist r0 r))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:forall x1 : R, pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i) -> g x1 = x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)

(pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2 < pos_Rl (cons_ORlist lf lg) (S i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:forall x1 : R, pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i) -> g x1 = x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)
H15:pos_Rl (cons_ORlist lf lg) i < (pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2 < pos_Rl (cons_ORlist lf lg) (S i)
x0 = g ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i = pos_Rl (cons_ORlist lf lg) (S i)
g x = g ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Rlength (mid_Rlist r0 r))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:forall x1 : R, pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i) -> g x1 = x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)
H15:pos_Rl (cons_ORlist lf lg) i < (pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2 < pos_Rl (cons_ORlist lf lg) (S i)

x0 = g ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i = pos_Rl (cons_ORlist lf lg) (S i)
g x = g ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Rlength (mid_Rlist r0 r))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
H13:pos_Rl (cons_ORlist lf lg) i <= pos_Rl (cons_ORlist lf lg) (S i)
H14:pos_Rl (cons_ORlist lf lg) i = pos_Rl (cons_ORlist lf lg) (S i)

g x = g ((pos_Rl (cons_ORlist lf lg) i + pos_Rl (cons_ORlist lf lg) (S i)) / 2)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Rlength (mid_Rlist r0 r))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0

(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0
(i < Rlength (mid_Rlist r0 r))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
x0:R
H11, H12:constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) x0
r:R
r0:Rlist
Hyp_cons:cons_ORlist lf lg = cons r r0

(i < Rlength (mid_Rlist r0 r))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

a < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

a <= pos_Rl (cons_ORlist lf lg) i
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl (cons_ORlist lf lg) i < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

pos_Rl (cons_ORlist lf lg) 0 <= pos_Rl (cons_ORlist lf lg) i
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl lf 0 = pos_Rl lg 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl (cons_ORlist lf lg) i < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

ordered_Rlist (cons_ORlist lf lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(0 <= i)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(i < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl lf 0 = pos_Rl lg 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl (cons_ORlist lf lg) i < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

(0 <= i)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(i < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl lf 0 = pos_Rl lg 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl (cons_ORlist lf lg) i < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

(i < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl lf 0 = pos_Rl lg 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl (cons_ORlist lf lg) i < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

pos_Rl lf 0 = pos_Rl lg 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl (cons_ORlist lf lg) i < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

pos_Rl (cons_ORlist lf lg) i < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl (cons_ORlist lf lg) (S i) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

pos_Rl (cons_ORlist lf lg) (S i) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl lf (Init.Nat.pred (Rlength lf)) = pos_Rl lg (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

ordered_Rlist (cons_ORlist lf lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(S i <= Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(Init.Nat.pred (Rlength (cons_ORlist lf lg)) < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl lf (Init.Nat.pred (Rlength lf)) = pos_Rl lg (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

(S i <= Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(Init.Nat.pred (Rlength (cons_ORlist lf lg)) < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl lf (Init.Nat.pred (Rlength lf)) = pos_Rl lg (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H12:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

(Init.Nat.pred (Rlength (cons_ORlist lf lg)) < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
pos_Rl lf (Init.Nat.pred (Rlength lf)) = pos_Rl lg (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)

pos_Rl lf (Init.Nat.pred (Rlength lf)) = pos_Rl lg (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b

exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop

Nbound I
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I

exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I

exists n : nat, I n
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I

pos_Rl lg 0 <= pos_Rl (cons_ORlist lf lg) i
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
(0 < Rlength lg)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I

pos_Rl lg 0 <= pos_Rl (cons_ORlist lf lg) 0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
pos_Rl (cons_ORlist lf lg) 0 <= pos_Rl (cons_ORlist lf lg) i
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
(0 < Rlength lg)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I

pos_Rl (cons_ORlist lf lg) 0 <= pos_Rl (cons_ORlist lf lg) i
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
(0 < Rlength lg)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I

(0 < Rlength lg)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n

exists l : R, constant_D_eq g (open_interval (pos_Rl (cons_ORlist lf lg) i) (pos_Rl (cons_ORlist lf lg) (S i))) l
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)

(x0 < Init.Nat.pred (Rlength lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
g x1 = pos_Rl lg0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat

(S x0 < Rlength lg)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
Rlength lg = S (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
g x1 = pos_Rl lg0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg

(S x0 < S x0)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
m:nat
H20:(S x0 <= m)%nat
H19:S m = Rlength lg
(S x0 < S m)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
Rlength lg = S (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
g x1 = pos_Rl lg0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg

(S x0 < S x0)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
Rlength lg = S (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
g x1 = pos_Rl lg0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg

x0 = Init.Nat.pred (Rlength lg) -> (S x0 < S x0)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
x0 = Init.Nat.pred (Rlength lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
Rlength lg = S (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
g x1 = pos_Rl lg0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
H19:x0 = Init.Nat.pred (Rlength lg)

pos_Rl (cons_ORlist lf lg) i < b -> (S x0 < S x0)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
H19:x0 = Init.Nat.pred (Rlength lg)
pos_Rl (cons_ORlist lf lg) i < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
x0 = Init.Nat.pred (Rlength lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
Rlength lg = S (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
g x1 = pos_Rl lg0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
H19:x0 = Init.Nat.pred (Rlength lg)

pos_Rl (cons_ORlist lf lg) i < b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
x0 = Init.Nat.pred (Rlength lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
Rlength lg = S (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
g x1 = pos_Rl lg0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
H19:x0 = Init.Nat.pred (Rlength lg)

pos_Rl (cons_ORlist lf lg) i < pos_Rl (cons_ORlist lf lg) (S i)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
H19:x0 = Init.Nat.pred (Rlength lg)
pos_Rl (cons_ORlist lf lg) (S i) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
x0 = Init.Nat.pred (Rlength lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
Rlength lg = S (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
g x1 = pos_Rl lg0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
H19:x0 = Init.Nat.pred (Rlength lg)

pos_Rl (cons_ORlist lf lg) (S i) <= b
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
x0 = Init.Nat.pred (Rlength lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
Rlength lg = S (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
g x1 = pos_Rl lg0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
H19:x0 = Init.Nat.pred (Rlength lg)

pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg)))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
H19:x0 = Init.Nat.pred (Rlength lg)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= pos_Rl lg (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
x0 = Init.Nat.pred (Rlength lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
Rlength lg = S (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
g x1 = pos_Rl lg0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
H19:x0 = Init.Nat.pred (Rlength lg)
H21:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H22:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

ordered_Rlist (cons_ORlist lf lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
H19:x0 = Init.Nat.pred (Rlength lg)
H21:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H22:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(S i <= Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
H19:x0 = Init.Nat.pred (Rlength lg)
H21:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H22:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(Init.Nat.pred (Rlength (cons_ORlist lf lg)) < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
H19:x0 = Init.Nat.pred (Rlength lg)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= pos_Rl lg (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
x0 = Init.Nat.pred (Rlength lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
Rlength lg = S (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
g x1 = pos_Rl lg0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
H19:x0 = Init.Nat.pred (Rlength lg)
H21:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H22:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

(S i <= Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
H19:x0 = Init.Nat.pred (Rlength lg)
H21:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H22:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)
(Init.Nat.pred (Rlength (cons_ORlist lf lg)) < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
H19:x0 = Init.Nat.pred (Rlength lg)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= pos_Rl lg (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
x0 = Init.Nat.pred (Rlength lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
Rlength lg = S (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
g x1 = pos_Rl lg0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
H19:x0 = Init.Nat.pred (Rlength lg)
H21:ordered_Rlist (cons_ORlist lf lg) -> forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j
H22:(forall i0 j : nat, (i0 <= j)%nat -> (j < Rlength (cons_ORlist lf lg))%nat -> pos_Rl (cons_ORlist lf lg) i0 <= pos_Rl (cons_ORlist lf lg) j) -> ordered_Rlist (cons_ORlist lf lg)

(Init.Nat.pred (Rlength (cons_ORlist lf lg)) < Rlength (cons_ORlist lf lg))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
H19:x0 = Init.Nat.pred (Rlength lg)
pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= pos_Rl lg (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
x0 = Init.Nat.pred (Rlength lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
Rlength lg = S (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
g x1 = pos_Rl lg0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
H19:x0 = Init.Nat.pred (Rlength lg)

pos_Rl (cons_ORlist lf lg) (Init.Nat.pred (Rlength (cons_ORlist lf lg))) <= pos_Rl lg (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
x0 = Init.Nat.pred (Rlength lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
Rlength lg = S (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
g x1 = pos_Rl lg0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:b <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
H19:x0 = Init.Nat.pred (Rlength lg)

pos_Rl lf (Init.Nat.pred (Rlength lf)) = pos_Rl lg (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg
x0 = Init.Nat.pred (Rlength lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
Rlength lg = S (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
g x1 = pos_Rl lg0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
H20:S x0 = Rlength lg

x0 = Init.Nat.pred (Rlength lg)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat
Rlength lg = S (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
g x1 = pos_Rl lg0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H18:(x0 < Rlength lg)%nat

Rlength lg = S (Init.Nat.pred (Rlength lg))
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
g x1 = pos_Rl lg0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat

g x1 = pos_Rl lg0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
H18:forall x2 : R, pos_Rl lg x0 < x2 < pos_Rl lg (S x0) -> g x2 = pos_Rl lg0 x0

pos_Rl lg0 x0 = pos_Rl lg0 x0
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
H18:forall x2 : R, pos_Rl lg x0 < x2 < pos_Rl lg (S x0) -> g x2 = pos_Rl lg0 x0
pos_Rl lg x0 < x1 < pos_Rl lg (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
H14:I x0 /\ (forall i0 : nat, I i0 -> (i0 <= x0)%nat)
x1:R
H15:pos_Rl (cons_ORlist lf lg) i < x1 < pos_Rl (cons_ORlist lf lg) (S i)
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
H18:forall x2 : R, pos_Rl lg x0 < x2 < pos_Rl lg (S x0) -> g x2 = pos_Rl lg0 x0

pos_Rl lg x0 < x1 < pos_Rl lg (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
H18:forall x2 : R, pos_Rl lg x0 < x2 < pos_Rl lg (S x0) -> g x2 = pos_Rl lg0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lg)%nat

pos_Rl lg x0 < x1
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
H18:forall x2 : R, pos_Rl lg x0 < x2 < pos_Rl lg (S x0) -> g x2 = pos_Rl lg0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lg)%nat
x1 < pos_Rl lg (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
H18:forall x2 : R, pos_Rl lg x0 < x2 < pos_Rl lg (S x0) -> g x2 = pos_Rl lg0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lg)%nat

x1 < pos_Rl lg (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
H18:forall x2 : R, pos_Rl lg x0 < x2 < pos_Rl lg (S x0) -> g x2 = pos_Rl lg0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lg)%nat

pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lg (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
H18:forall x2 : R, pos_Rl lg x0 < x2 < pos_Rl lg (S x0) -> g x2 = pos_Rl lg0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lg)%nat

(S x0 < Rlength lg)%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
H18:forall x2 : R, pos_Rl lg x0 < x2 < pos_Rl lg (S x0) -> g x2 = pos_Rl lg0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lg)%nat
H22:(S x0 < Rlength lg)%nat
pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lg (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
H18:forall x2 : R, pos_Rl lg x0 < x2 < pos_Rl lg (S x0) -> g x2 = pos_Rl lg0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lg)%nat

(S x0 < S (Init.Nat.pred (Rlength lg)))%nat
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
H18:forall x2 : R, pos_Rl lg x0 < x2 < pos_Rl lg (S x0) -> g x2 = pos_Rl lg0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lg)%nat
S (Init.Nat.pred (Rlength lg)) = Rlength lg
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
H18:forall x2 : R, pos_Rl lg x0 < x2 < pos_Rl lg (S x0) -> g x2 = pos_Rl lg0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lg)%nat
H22:(S x0 < Rlength lg)%nat
pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lg (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
H18:forall x2 : R, pos_Rl lg x0 < x2 < pos_Rl lg (S x0) -> g x2 = pos_Rl lg0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lg)%nat

S (Init.Nat.pred (Rlength lg)) = Rlength lg
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
H18:forall x2 : R, pos_Rl lg x0 < x2 < pos_Rl lg (S x0) -> g x2 = pos_Rl lg0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lg)%nat
H22:(S x0 < Rlength lg)%nat
pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lg (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
H18:forall x2 : R, pos_Rl lg x0 < x2 < pos_Rl lg (S x0) -> g x2 = pos_Rl lg0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lg)%nat
H22:(S x0 < Rlength lg)%nat

pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lg (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
H18:forall x2 : R, pos_Rl lg x0 < x2 < pos_Rl lg (S x0) -> g x2 = pos_Rl lg0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lg)%nat
H22:(S x0 < Rlength lg)%nat
a0:pos_Rl lg (S x0) <= pos_Rl (cons_ORlist lf lg) i

pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lg (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
H18:forall x2 : R, pos_Rl lg x0 < x2 < pos_Rl lg (S x0) -> g x2 = pos_Rl lg0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lg)%nat
H22:(S x0 < Rlength lg)%nat
a0:~ pos_Rl lg (S x0) <= pos_Rl (cons_ORlist lf lg) i
pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lg (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
H18:forall x2 : R, pos_Rl lg x0 < x2 < pos_Rl lg (S x0) -> g x2 = pos_Rl lg0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lg)%nat
H22:(S x0 < Rlength lg)%nat
a0:~ pos_Rl lg (S x0) <= pos_Rl (cons_ORlist lf lg) i

pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lg (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
H18:forall x2 : R, pos_Rl lg x0 < x2 < pos_Rl lg (S x0) -> g x2 = pos_Rl lg0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lg)%nat
H22:(S x0 < Rlength lg)%nat
a0:~ pos_Rl lg (S x0) <= pos_Rl (cons_ORlist lf lg) i

pos_Rl (cons_ORlist lf lg) i < pos_Rl lg (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
H18:forall x2 : R, pos_Rl lg x0 < x2 < pos_Rl lg (S x0) -> g x2 = pos_Rl lg0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lg)%nat
H22:(S x0 < Rlength lg)%nat
a0:~ pos_Rl lg (S x0) <= pos_Rl (cons_ORlist lf lg) i
H23:pos_Rl (cons_ORlist lf lg) i < pos_Rl lg (S x0)
pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lg (S x0)
a, b:R
f, g:R -> R
lf, lg:Rlist
Hyp:a <= b
lg0, lf0:Rlist
Hyp_min:Rmin a b = a
Hyp_max:Rmax a b = b
H:ordered_Rlist lg
H1:pos_Rl lg 0 = a
H0:pos_Rl lg (Init.Nat.pred (Rlength lg)) = b
H2:Rlength lg = S (Rlength lg0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg0 i0)
H3:ordered_Rlist lf
H6:pos_Rl lf 0 = a
H5:pos_Rl lf (Init.Nat.pred (Rlength lf)) = b
H7:Rlength lf = S (Rlength lf0)
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lf))%nat -> constant_D_eq f (open_interval (pos_Rl lf i0) (pos_Rl lf (S i0))) (pos_Rl lf0 i0)
i:nat
H8:(i < Init.Nat.pred (Rlength (cons_ORlist lf lg)))%nat
x:R
H10:pos_Rl (cons_ORlist lf lg) i < x < pos_Rl (cons_ORlist lf lg) (S i)
H11:a < b
I:=fun j : nat => pos_Rl lg j <= pos_Rl (cons_ORlist lf lg) i /\ (j < Rlength lg)%nat:nat -> Prop
H12:Nbound I
H13:exists n : nat, I n
x0:nat
x1:R
H16:(x0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg x0) (pos_Rl lg (S x0))) (pos_Rl lg0 x0)
H17:(x0 < Init.Nat.pred (Rlength lg))%nat
H18:forall x2 : R, pos_Rl lg x0 < x2 < pos_Rl lg (S x0) -> g x2 = pos_Rl lg0 x0
H15:pos_Rl (cons_ORlist lf lg) i < x1
H19:x1 < pos_Rl (cons_ORlist lf lg) (S i)
H20:forall i0 : nat, I i0 -> (i0 <= x0)%nat
H14:pos_Rl lg x0 <= pos_Rl (cons_ORlist lf lg) i
H21:(x0 < Rlength lg)%nat
H22:(S x0 < Rlength lg)%nat
a0:~ pos_Rl lg (S x0) <= pos_Rl (cons_ORlist lf lg) i
H23:pos_Rl (cons_ORlist lf lg) i < pos_Rl lg (S x0)

pos_Rl (cons_ORlist lf lg) (S i) <= pos_Rl lg (S x0)
clear a0; apply RList_P17; try assumption; [ apply RList_P2; assumption | elim (RList_P9 lf lg (pos_Rl lg (S x0))); intros; apply H25; right; elim (RList_P3 lg (pos_Rl lg (S x0))); intros; apply H27; exists (S x0); split; [ reflexivity | apply H22 ] ]. Qed.

forall (a b : R) (f g : R -> R) (lf lg : Rlist), is_subdivision f a b lf -> is_subdivision g a b lg -> is_subdivision g a b (cons_ORlist lf lg)

forall (a b : R) (f g : R -> R) (lf lg : Rlist), is_subdivision f a b lf -> is_subdivision g a b lg -> is_subdivision g a b (cons_ORlist lf lg)
intros a b f g lf lg H H0; case (Rle_dec a b); intro; [ apply StepFun_P24 with f; assumption | apply StepFun_P5; apply StepFun_P24 with f; [ auto with real | apply StepFun_P5; assumption | apply StepFun_P5; assumption ] ]. Qed.

forall (a b l : R) (f g : R -> R) (l1 : Rlist), is_subdivision f a b l1 -> is_subdivision g a b l1 -> is_subdivision (fun x : R => f x + l * g x) a b l1

forall (a b l : R) (f g : R -> R) (l1 : Rlist), is_subdivision f a b l1 -> is_subdivision g a b l1 -> is_subdivision (fun x : R => f x + l * g x) a b l1
a, b, l:R
f, g:R -> R
l1, x0:Rlist
H0:ordered_Rlist l1
H1:pos_Rl l1 0 = Rmin a b
H2:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H3:Rlength l1 = S (Rlength x0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl x0 i)
x:Rlist
H9:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq g (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl x i)

is_subdivision (fun x1 : R => f x1 + l * g x1) a b l1
a, b, l:R
f, g:R -> R
l1, x0:Rlist
H0:ordered_Rlist l1
H1:pos_Rl l1 0 = Rmin a b
H2:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H3:Rlength l1 = S (Rlength x0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl x0 i)
x:Rlist
H9:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq g (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl x i)

Rlength l1 = S (Rlength (FF l1 (fun x1 : R => f x1 + l * g x1)))
a, b, l:R
f, g:R -> R
l1, x0:Rlist
H0:ordered_Rlist l1
H1:pos_Rl l1 0 = Rmin a b
H2:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H3:Rlength l1 = S (Rlength x0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl x0 i)
x:Rlist
H9:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq g (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl x i)
forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq (fun x1 : R => f x1 + l * g x1) (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl (FF l1 (fun x1 : R => f x1 + l * g x1)) i)
a, b, l:R
f, g:R -> R
l1, x0:Rlist
H0:ordered_Rlist l1
H1:pos_Rl l1 0 = Rmin a b
H2:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H3:Rlength l1 = S (Rlength x0)
H4:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl x0 i)
x:Rlist
H9:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq g (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl x i)

forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq (fun x1 : R => f x1 + l * g x1) (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl (FF l1 (fun x1 : R => f x1 + l * g x1)) i)
a, b, l:R
f, g:R -> R
l1, x0:Rlist
H0:ordered_Rlist l1
H1:pos_Rl l1 0 = Rmin a b
H2:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H3:Rlength l1 = S (Rlength x0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x0 i0)
x:Rlist
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq g (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x i0)
i:nat
H8:(i < Init.Nat.pred (Rlength l1))%nat
x1:R
H10:pos_Rl l1 i < x1 < pos_Rl l1 (S i)

l1 <> nil
a, b, l:R
f, g:R -> R
l1, x0:Rlist
H0:ordered_Rlist l1
H1:pos_Rl l1 0 = Rmin a b
H2:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H3:Rlength l1 = S (Rlength x0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x0 i0)
x:Rlist
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq g (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x i0)
i:nat
H8:(i < Init.Nat.pred (Rlength l1))%nat
x1:R
H10:pos_Rl l1 i < x1 < pos_Rl l1 (S i)
H11:l1 <> nil
pos_Rl x0 i + l * pos_Rl x i = pos_Rl (FF l1 (fun x2 : R => f x2 + l * g x2)) i
a, b, l:R
f, g:R -> R
l1, x0:Rlist
H0:ordered_Rlist l1
H1:pos_Rl l1 0 = Rmin a b
H2:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H3:Rlength l1 = S (Rlength x0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x0 i0)
x:Rlist
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq g (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x i0)
i:nat
H8:(i < Init.Nat.pred (Rlength l1))%nat
x1:R
H10:pos_Rl l1 i < x1 < pos_Rl l1 (S i)
H11:l1 <> nil

pos_Rl x0 i + l * pos_Rl x i = pos_Rl (FF l1 (fun x2 : R => f x2 + l * g x2)) i
a, b, l:R
f, g:R -> R
l1, x0:Rlist
H0:ordered_Rlist l1
H1:pos_Rl l1 0 = Rmin a b
H2:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H3:Rlength l1 = S (Rlength x0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x0 i0)
x:Rlist
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq g (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x i0)
i:nat
H8:(i < Init.Nat.pred (Rlength l1))%nat
x1:R
H10:pos_Rl l1 i < x1 < pos_Rl l1 (S i)
H11:l1 <> nil
r:R
r0:Rlist
H12:l1 = cons r r0

pos_Rl x0 i + l * pos_Rl x i = f (pos_Rl (mid_Rlist (cons r r0) r) (S i)) + l * g (pos_Rl (mid_Rlist (cons r r0) r) (S i))
a, b, l:R
f, g:R -> R
l1, x0:Rlist
H0:ordered_Rlist l1
H1:pos_Rl l1 0 = Rmin a b
H2:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H3:Rlength l1 = S (Rlength x0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x0 i0)
x:Rlist
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq g (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x i0)
i:nat
H8:(i < Init.Nat.pred (Rlength l1))%nat
x1:R
H10:pos_Rl l1 i < x1 < pos_Rl l1 (S i)
H11:l1 <> nil
r:R
r0:Rlist
H12:l1 = cons r r0
(S i < Rlength (mid_Rlist (cons r r0) r))%nat
a, b, l:R
f, g:R -> R
l1, x0:Rlist
H0:ordered_Rlist l1
H1:pos_Rl l1 0 = Rmin a b
H2:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H3:Rlength l1 = S (Rlength x0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x0 i0)
x:Rlist
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq g (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x i0)
i:nat
H8:(i < Init.Nat.pred (Rlength l1))%nat
x1:R
H10:pos_Rl l1 i < x1 < pos_Rl l1 (S i)
H11:l1 <> nil
r:R
r0:Rlist
H12:l1 = cons r r0

pos_Rl x0 i + l * pos_Rl x i = f ((pos_Rl (cons r r0) i + pos_Rl (cons r r0) (S i)) / 2) + l * g ((pos_Rl (cons r r0) i + pos_Rl (cons r r0) (S i)) / 2)
a, b, l:R
f, g:R -> R
l1, x0:Rlist
H0:ordered_Rlist l1
H1:pos_Rl l1 0 = Rmin a b
H2:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H3:Rlength l1 = S (Rlength x0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x0 i0)
x:Rlist
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq g (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x i0)
i:nat
H8:(i < Init.Nat.pred (Rlength l1))%nat
x1:R
H10:pos_Rl l1 i < x1 < pos_Rl l1 (S i)
H11:l1 <> nil
r:R
r0:Rlist
H12:l1 = cons r r0
(i < Init.Nat.pred (Rlength (cons r r0)))%nat
a, b, l:R
f, g:R -> R
l1, x0:Rlist
H0:ordered_Rlist l1
H1:pos_Rl l1 0 = Rmin a b
H2:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H3:Rlength l1 = S (Rlength x0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x0 i0)
x:Rlist
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq g (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x i0)
i:nat
H8:(i < Init.Nat.pred (Rlength l1))%nat
x1:R
H10:pos_Rl l1 i < x1 < pos_Rl l1 (S i)
H11:l1 <> nil
r:R
r0:Rlist
H12:l1 = cons r r0
(S i < Rlength (mid_Rlist (cons r r0) r))%nat
a, b, l:R
f, g:R -> R
l1, x0:Rlist
H0:ordered_Rlist l1
H1:pos_Rl l1 0 = Rmin a b
H2:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H3:Rlength l1 = S (Rlength x0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x0 i0)
x:Rlist
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq g (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x i0)
i:nat
H8:(i < Init.Nat.pred (Rlength l1))%nat
x1:R
H10:pos_Rl l1 i < x1 < pos_Rl l1 (S i)
H11:l1 <> nil
r:R
r0:Rlist
H12:l1 = cons r r0

(i < Init.Nat.pred (Rlength (cons r r0)))%nat
a, b, l:R
f, g:R -> R
l1, x0:Rlist
H0:ordered_Rlist l1
H1:pos_Rl l1 0 = Rmin a b
H2:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H3:Rlength l1 = S (Rlength x0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x0 i0)
x:Rlist
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq g (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x i0)
i:nat
H8:(i < Init.Nat.pred (Rlength l1))%nat
x1:R
H10:pos_Rl l1 i < x1 < pos_Rl l1 (S i)
H11:l1 <> nil
r:R
r0:Rlist
H12:l1 = cons r r0
(S i < Rlength (mid_Rlist (cons r r0) r))%nat
a, b, l:R
f, g:R -> R
l1, x0:Rlist
H0:ordered_Rlist l1
H1:pos_Rl l1 0 = Rmin a b
H2:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H3:Rlength l1 = S (Rlength x0)
H4:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x0 i0)
x:Rlist
H9:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq g (fun x2 : R => pos_Rl l1 i0 < x2 < pos_Rl l1 (S i0)) (pos_Rl x i0)
i:nat
H8:(i < Init.Nat.pred (Rlength l1))%nat
x1:R
H10:pos_Rl l1 i < x1 < pos_Rl l1 (S i)
H11:l1 <> nil
r:R
r0:Rlist
H12:l1 = cons r r0

(S i < Rlength (mid_Rlist (cons r r0) r))%nat
rewrite RList_P14; simpl; rewrite H12 in H8; simpl in H8; apply lt_n_S; apply H8. Qed.

forall (a b l : R) (f g : R -> R) (lf lg : Rlist), is_subdivision f a b lf -> is_subdivision g a b lg -> is_subdivision (fun x : R => f x + l * g x) a b (cons_ORlist lf lg)

forall (a b l : R) (f g : R -> R) (lf lg : Rlist), is_subdivision f a b lf -> is_subdivision g a b lg -> is_subdivision (fun x : R => f x + l * g x) a b (cons_ORlist lf lg)
intros a b l f g lf lg H H0; apply StepFun_P26; [ apply StepFun_P23 with g; assumption | apply StepFun_P25 with f; assumption ]. Qed.
The set of step functions on a,b is a vectorial space

forall (a b l : R) (f g : StepFun a b), IsStepFun (fun x : R => f x + l * g x) a b

forall (a b l : R) (f g : StepFun a b), IsStepFun (fun x : R => f x + l * g x) a b
intros a b l f g; unfold IsStepFun; assert (H := pre f); assert (H0 := pre g); unfold IsStepFun in H, H0; elim H; elim H0; intros; apply existT with (cons_ORlist x0 x); apply StepFun_P27; assumption. Qed.

forall (a b : R) (f : StepFun a b), is_subdivision f a b (subdivision f)

forall (a b : R) (f : StepFun a b), is_subdivision f a b (subdivision f)
intros a b f; unfold is_subdivision; apply existT with (subdivision_val f); apply StepFun_P1. Qed.

forall (a b l : R) (f g : StepFun a b), RiemannInt_SF {| fe := fun x : R => f x + l * g x; pre := StepFun_P28 l f g |} = RiemannInt_SF f + l * RiemannInt_SF g

forall (a b l : R) (f g : StepFun a b), RiemannInt_SF {| fe := fun x : R => f x + l * g x; pre := StepFun_P28 l f g |} = RiemannInt_SF f + l * RiemannInt_SF g
intros a b l f g; unfold RiemannInt_SF; case (Rle_dec a b); (intro; replace (Int_SF (subdivision_val (mkStepFun (StepFun_P28 l f g))) (subdivision (mkStepFun (StepFun_P28 l f g)))) with (Int_SF (FF (cons_ORlist (subdivision f) (subdivision g)) (fun x:R => f x + l * g x)) (cons_ORlist (subdivision f) (subdivision g))); [ rewrite StepFun_P19; replace (Int_SF (FF (cons_ORlist (subdivision f) (subdivision g)) f) (cons_ORlist (subdivision f) (subdivision g))) with (Int_SF (subdivision_val f) (subdivision f)); [ replace (Int_SF (FF (cons_ORlist (subdivision f) (subdivision g)) g) (cons_ORlist (subdivision f) (subdivision g))) with (Int_SF (subdivision_val g) (subdivision g)); [ ring | apply StepFun_P17 with (fe g) a b; [ apply StepFun_P1 | apply StepFun_P21; apply StepFun_P25 with (fe f); apply StepFun_P29 ] ] | apply StepFun_P17 with (fe f) a b; [ apply StepFun_P1 | apply StepFun_P21; apply StepFun_P23 with (fe g); apply StepFun_P29 ] ] | apply StepFun_P17 with (fun x:R => f x + l * g x) a b; [ apply StepFun_P21; apply StepFun_P27; apply StepFun_P29 | apply (StepFun_P1 (mkStepFun (StepFun_P28 l f g))) ] ]). Qed.

forall (a b : R) (f : R -> R) (l lf : Rlist), adapted_couple f a b l lf -> adapted_couple (fun x : R => Rabs (f x)) a b l (app_Rlist lf Rabs)

forall (a b : R) (f : R -> R) (l lf : Rlist), adapted_couple f a b l lf -> adapted_couple (fun x : R => Rabs (f x)) a b l (app_Rlist lf Rabs)
a, b:R
f:R -> R
l, lf:Rlist
H0:ordered_Rlist l
H2:pos_Rl l 0 = Rmin a b
H1:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H3:Rlength l = S (Rlength lf)
H5:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)

Rlength l = S (Rlength (app_Rlist lf Rabs))
a, b:R
f:R -> R
l, lf:Rlist
H0:ordered_Rlist l
H2:pos_Rl l 0 = Rmin a b
H1:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H3:Rlength l = S (Rlength lf)
H5:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)
forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq (fun x : R => Rabs (f x)) (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl (app_Rlist lf Rabs) i)
a, b:R
f:R -> R
l, lf:Rlist
H0:ordered_Rlist l
H2:pos_Rl l 0 = Rmin a b
H1:pos_Rl l (Init.Nat.pred (Rlength l)) = Rmax a b
H3:Rlength l = S (Rlength lf)
H5:forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq f (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl lf i)

forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq (fun x : R => Rabs (f x)) (open_interval (pos_Rl l i) (pos_Rl l (S i))) (pos_Rl (app_Rlist lf Rabs) i)
intros; unfold constant_D_eq, open_interval; unfold constant_D_eq, open_interval in H5; intros; rewrite (H5 _ H _ H4); rewrite RList_P12; [ reflexivity | rewrite H3 in H; simpl in H; apply H ]. Qed.

forall (a b : R) (f : StepFun a b), IsStepFun (fun x : R => Rabs (f x)) a b

forall (a b : R) (f : StepFun a b), IsStepFun (fun x : R => Rabs (f x)) a b
intros a b f; unfold IsStepFun; apply existT with (subdivision f); unfold is_subdivision; apply existT with (app_Rlist (subdivision_val f) Rabs); apply StepFun_P31; apply StepFun_P1. Qed.

forall l2 l1 : Rlist, ordered_Rlist l1 -> Rabs (Int_SF l2 l1) <= Int_SF (app_Rlist l2 Rabs) l1

forall l2 l1 : Rlist, ordered_Rlist l1 -> Rabs (Int_SF l2 l1) <= Int_SF (app_Rlist l2 Rabs) l1
l2, l1:Rlist
H:ordered_Rlist l1

Rabs (Int_SF nil l1) <= Int_SF (app_Rlist nil Rabs) l1
l2:Rlist
r:R
r0:Rlist
H:forall l0 : Rlist, ordered_Rlist l0 -> Rabs (Int_SF r0 l0) <= Int_SF (app_Rlist r0 Rabs) l0
l1:Rlist
H0:ordered_Rlist l1
Rabs (Int_SF (cons r r0) l1) <= Int_SF (app_Rlist (cons r r0) Rabs) l1
l2:Rlist
r:R
r0:Rlist
H:forall l0 : Rlist, ordered_Rlist l0 -> Rabs (Int_SF r0 l0) <= Int_SF (app_Rlist r0 Rabs) l0
l1:Rlist
H0:ordered_Rlist l1

Rabs (Int_SF (cons r r0) l1) <= Int_SF (app_Rlist (cons r r0) Rabs) l1
l2:Rlist
r:R
r0:Rlist
H:forall l1 : Rlist, ordered_Rlist l1 -> Rabs (Int_SF r0 l1) <= Int_SF (app_Rlist r0 Rabs) l1
H0:ordered_Rlist nil

Rabs 0 <= 0
l2:Rlist
r:R
r0:Rlist
H:forall l0 : Rlist, ordered_Rlist l0 -> Rabs (Int_SF r0 l0) <= Int_SF (app_Rlist r0 Rabs) l0
r1:R
l1:Rlist
H0:ordered_Rlist (cons r1 l1)
Hrecl1:ordered_Rlist l1 -> Rabs match l1 with | nil => 0 | cons x nil => 0 | cons x (cons y k') => r * (y - x) + Int_SF r0 (cons y k') end <= match l1 with | nil => 0 | cons x nil => 0 | cons x (cons y k') => Rabs r * (y - x) + Int_SF (app_Rlist r0 Rabs) (cons y k') end
Rabs match l1 with | nil => 0 | cons y k' => r * (y - r1) + Int_SF r0 (cons y k') end <= match l1 with | nil => 0 | cons y k' => Rabs r * (y - r1) + Int_SF (app_Rlist r0 Rabs) (cons y k') end
l2:Rlist
r:R
r0:Rlist
H:forall l0 : Rlist, ordered_Rlist l0 -> Rabs (Int_SF r0 l0) <= Int_SF (app_Rlist r0 Rabs) l0
r1:R
l1:Rlist
H0:ordered_Rlist (cons r1 l1)
Hrecl1:ordered_Rlist l1 -> Rabs match l1 with | nil => 0 | cons x nil => 0 | cons x (cons y k') => r * (y - x) + Int_SF r0 (cons y k') end <= match l1 with | nil => 0 | cons x nil => 0 | cons x (cons y k') => Rabs r * (y - x) + Int_SF (app_Rlist r0 Rabs) (cons y k') end

Rabs match l1 with | nil => 0 | cons y k' => r * (y - r1) + Int_SF r0 (cons y k') end <= match l1 with | nil => 0 | cons y k' => Rabs r * (y - r1) + Int_SF (app_Rlist r0 Rabs) (cons y k') end
l2:Rlist
r:R
r0:Rlist
H:forall l1 : Rlist, ordered_Rlist l1 -> Rabs (Int_SF r0 l1) <= Int_SF (app_Rlist r0 Rabs) l1
r1:R
H0:ordered_Rlist (cons r1 nil)
Hrecl1:ordered_Rlist nil -> Rabs 0 <= 0

Rabs 0 <= 0
l2:Rlist
r:R
r0:Rlist
H:forall l0 : Rlist, ordered_Rlist l0 -> Rabs (Int_SF r0 l0) <= Int_SF (app_Rlist r0 Rabs) l0
r1, r2:R
l1:Rlist
H0:ordered_Rlist (cons r1 (cons r2 l1))
Hrecl1:ordered_Rlist (cons r2 l1) -> Rabs match l1 with | nil => 0 | cons y k' => r * (y - r2) + Int_SF r0 (cons y k') end <= match l1 with | nil => 0 | cons y k' => Rabs r * (y - r2) + Int_SF (app_Rlist r0 Rabs) (cons y k') end
Hrecl0:ordered_Rlist (cons r1 l1) -> (ordered_Rlist l1 -> Rabs match l1 with | nil => 0 | cons x nil => 0 | cons x (cons y k') => r * (y - x) + Int_SF r0 (cons y k') end <= match l1 with | nil => 0 | cons x nil => 0 | cons x (cons y k') => Rabs r * (y - x) + Int_SF (app_Rlist r0 Rabs) (cons y k') end) -> Rabs match l1 with | nil => 0 | cons y k' => r * (y - r1) + Int_SF r0 (cons y k') end <= match l1 with | nil => 0 | cons y k' => Rabs r * (y - r1) + Int_SF (app_Rlist r0 Rabs) (cons y k') end
Rabs (r * (r2 - r1) + Int_SF r0 (cons r2 l1)) <= Rabs r * (r2 - r1) + Int_SF (app_Rlist r0 Rabs) (cons r2 l1)
l2:Rlist
r:R
r0:Rlist
H:forall l0 : Rlist, ordered_Rlist l0 -> Rabs (Int_SF r0 l0) <= Int_SF (app_Rlist r0 Rabs) l0
r1, r2:R
l1:Rlist
H0:ordered_Rlist (cons r1 (cons r2 l1))
Hrecl1:ordered_Rlist (cons r2 l1) -> Rabs match l1 with | nil => 0 | cons y k' => r * (y - r2) + Int_SF r0 (cons y k') end <= match l1 with | nil => 0 | cons y k' => Rabs r * (y - r2) + Int_SF (app_Rlist r0 Rabs) (cons y k') end
Hrecl0:ordered_Rlist (cons r1 l1) -> (ordered_Rlist l1 -> Rabs match l1 with | nil => 0 | cons x nil => 0 | cons x (cons y k') => r * (y - x) + Int_SF r0 (cons y k') end <= match l1 with | nil => 0 | cons x nil => 0 | cons x (cons y k') => Rabs r * (y - x) + Int_SF (app_Rlist r0 Rabs) (cons y k') end) -> Rabs match l1 with | nil => 0 | cons y k' => r * (y - r1) + Int_SF r0 (cons y k') end <= match l1 with | nil => 0 | cons y k' => Rabs r * (y - r1) + Int_SF (app_Rlist r0 Rabs) (cons y k') end

Rabs (r * (r2 - r1) + Int_SF r0 (cons r2 l1)) <= Rabs r * (r2 - r1) + Int_SF (app_Rlist r0 Rabs) (cons r2 l1)
l2:Rlist
r:R
r0:Rlist
H:forall l0 : Rlist, ordered_Rlist l0 -> Rabs (Int_SF r0 l0) <= Int_SF (app_Rlist r0 Rabs) l0
r1, r2:R
l1:Rlist
H0:ordered_Rlist (cons r1 (cons r2 l1))
Hrecl1:ordered_Rlist (cons r2 l1) -> Rabs match l1 with | nil => 0 | cons y k' => r * (y - r2) + Int_SF r0 (cons y k') end <= match l1 with | nil => 0 | cons y k' => Rabs r * (y - r2) + Int_SF (app_Rlist r0 Rabs) (cons y k') end
Hrecl0:ordered_Rlist (cons r1 l1) -> (ordered_Rlist l1 -> Rabs match l1 with | nil => 0 | cons x nil => 0 | cons x (cons y k') => r * (y - x) + Int_SF r0 (cons y k') end <= match l1 with | nil => 0 | cons x nil => 0 | cons x (cons y k') => Rabs r * (y - x) + Int_SF (app_Rlist r0 Rabs) (cons y k') end) -> Rabs match l1 with | nil => 0 | cons y k' => r * (y - r1) + Int_SF r0 (cons y k') end <= match l1 with | nil => 0 | cons y k' => Rabs r * (y - r1) + Int_SF (app_Rlist r0 Rabs) (cons y k') end

Rabs (r * (r2 - r1) + Int_SF r0 (cons r2 l1)) <= Rabs (r * (r2 - r1)) + Rabs (Int_SF r0 (cons r2 l1))
l2:Rlist
r:R
r0:Rlist
H:forall l0 : Rlist, ordered_Rlist l0 -> Rabs (Int_SF r0 l0) <= Int_SF (app_Rlist r0 Rabs) l0
r1, r2:R
l1:Rlist
H0:ordered_Rlist (cons r1 (cons r2 l1))
Hrecl1:ordered_Rlist (cons r2 l1) -> Rabs match l1 with | nil => 0 | cons y k' => r * (y - r2) + Int_SF r0 (cons y k') end <= match l1 with | nil => 0 | cons y k' => Rabs r * (y - r2) + Int_SF (app_Rlist r0 Rabs) (cons y k') end
Hrecl0:ordered_Rlist (cons r1 l1) -> (ordered_Rlist l1 -> Rabs match l1 with | nil => 0 | cons x nil => 0 | cons x (cons y k') => r * (y - x) + Int_SF r0 (cons y k') end <= match l1 with | nil => 0 | cons x nil => 0 | cons x (cons y k') => Rabs r * (y - x) + Int_SF (app_Rlist r0 Rabs) (cons y k') end) -> Rabs match l1 with | nil => 0 | cons y k' => r * (y - r1) + Int_SF r0 (cons y k') end <= match l1 with | nil => 0 | cons y k' => Rabs r * (y - r1) + Int_SF (app_Rlist r0 Rabs) (cons y k') end
Rabs (r * (r2 - r1)) + Rabs (Int_SF r0 (cons r2 l1)) <= Rabs r * (r2 - r1) + Int_SF (app_Rlist r0 Rabs) (cons r2 l1)
l2:Rlist
r:R
r0:Rlist
H:forall l0 : Rlist, ordered_Rlist l0 -> Rabs (Int_SF r0 l0) <= Int_SF (app_Rlist r0 Rabs) l0
r1, r2:R
l1:Rlist
H0:ordered_Rlist (cons r1 (cons r2 l1))
Hrecl1:ordered_Rlist (cons r2 l1) -> Rabs match l1 with | nil => 0 | cons y k' => r * (y - r2) + Int_SF r0 (cons y k') end <= match l1 with | nil => 0 | cons y k' => Rabs r * (y - r2) + Int_SF (app_Rlist r0 Rabs) (cons y k') end
Hrecl0:ordered_Rlist (cons r1 l1) -> (ordered_Rlist l1 -> Rabs match l1 with | nil => 0 | cons x nil => 0 | cons x (cons y k') => r * (y - x) + Int_SF r0 (cons y k') end <= match l1 with | nil => 0 | cons x nil => 0 | cons x (cons y k') => Rabs r * (y - x) + Int_SF (app_Rlist r0 Rabs) (cons y k') end) -> Rabs match l1 with | nil => 0 | cons y k' => r * (y - r1) + Int_SF r0 (cons y k') end <= match l1 with | nil => 0 | cons y k' => Rabs r * (y - r1) + Int_SF (app_Rlist r0 Rabs) (cons y k') end

Rabs (r * (r2 - r1)) + Rabs (Int_SF r0 (cons r2 l1)) <= Rabs r * (r2 - r1) + Int_SF (app_Rlist r0 Rabs) (cons r2 l1)
rewrite Rabs_mult; rewrite (Rabs_right (r2 - r1)); [ apply Rplus_le_compat_l; apply H; apply RList_P4 with r1; assumption | apply Rge_minus; apply Rle_ge; apply (H0 0%nat); simpl; apply lt_O_Sn ]. Qed.

forall (a b : R) (f : StepFun a b), a <= b -> Rabs (RiemannInt_SF f) <= RiemannInt_SF {| fe := fun x : R => Rabs (f x); pre := StepFun_P32 f |}

forall (a b : R) (f : StepFun a b), a <= b -> Rabs (RiemannInt_SF f) <= RiemannInt_SF {| fe := fun x : R => Rabs (f x); pre := StepFun_P32 f |}
a, b:R
f:StepFun a b
H:a <= b

Rabs (Int_SF (subdivision_val f) (subdivision f)) <= Int_SF (subdivision_val {| fe := fun x : R => Rabs (f x); pre := StepFun_P32 f |}) (subdivision {| fe := fun x : R => Rabs (f x); pre := StepFun_P32 f |})
a, b:R
f:StepFun a b
H:a <= b

Rabs (Int_SF (subdivision_val f) (subdivision f)) <= Int_SF (app_Rlist (subdivision_val f) Rabs) (subdivision f)
a, b:R
f:StepFun a b
H:a <= b
Int_SF (app_Rlist (subdivision_val f) Rabs) (subdivision f) = Int_SF (subdivision_val {| fe := fun x : R => Rabs (f x); pre := StepFun_P32 f |}) (subdivision {| fe := fun x : R => Rabs (f x); pre := StepFun_P32 f |})
a, b:R
f:StepFun a b
H:a <= b

Int_SF (app_Rlist (subdivision_val f) Rabs) (subdivision f) = Int_SF (subdivision_val {| fe := fun x : R => Rabs (f x); pre := StepFun_P32 f |}) (subdivision {| fe := fun x : R => Rabs (f x); pre := StepFun_P32 f |})
apply StepFun_P17 with (fun x:R => Rabs (f x)) a b; [ apply StepFun_P31; apply StepFun_P1 | apply (StepFun_P1 (mkStepFun (StepFun_P32 f))) ]. Qed.

forall (l : Rlist) (a b : R) (f g : R -> R), ordered_Rlist l -> pos_Rl l 0 = a -> pos_Rl l (Init.Nat.pred (Rlength l)) = b -> (forall x : R, a < x < b -> f x <= g x) -> Int_SF (FF l f) l <= Int_SF (FF l g) l

forall (l : Rlist) (a b : R) (f g : R -> R), ordered_Rlist l -> pos_Rl l 0 = a -> pos_Rl l (Init.Nat.pred (Rlength l)) = b -> (forall x : R, a < x < b -> f x <= g x) -> Int_SF (FF l f) l <= Int_SF (FF l g) l
l:Rlist
a, b:R
f, g:R -> R
H:ordered_Rlist nil
H0:pos_Rl nil 0 = a
H1:pos_Rl nil (Init.Nat.pred (Rlength nil)) = b
H2:forall x : R, a < x < b -> f x <= g x

Int_SF (FF nil f) nil <= Int_SF (FF nil g) nil
l:Rlist
r:R
r0:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r0 -> pos_Rl r0 0 = a0 -> pos_Rl r0 (Init.Nat.pred (Rlength r0)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r0 f0) r0 <= Int_SF (FF r0 g0) r0
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r r0)
H1:pos_Rl (cons r r0) 0 = a
H2:pos_Rl (cons r r0) (Init.Nat.pred (Rlength (cons r r0))) = b
H3:forall x : R, a < x < b -> f x <= g x
Int_SF (FF (cons r r0) f) (cons r r0) <= Int_SF (FF (cons r r0) g) (cons r r0)
l:Rlist
r:R
r0:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r0 -> pos_Rl r0 0 = a0 -> pos_Rl r0 (Init.Nat.pred (Rlength r0)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r0 f0) r0 <= Int_SF (FF r0 g0) r0
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r r0)
H1:pos_Rl (cons r r0) 0 = a
H2:pos_Rl (cons r r0) (Init.Nat.pred (Rlength (cons r r0))) = b
H3:forall x : R, a < x < b -> f x <= g x

Int_SF (FF (cons r r0) f) (cons r r0) <= Int_SF (FF (cons r r0) g) (cons r r0)
l:Rlist
r:R
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist nil -> pos_Rl nil 0 = a0 -> pos_Rl nil (Init.Nat.pred (Rlength nil)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF nil f0) nil <= Int_SF (FF nil g0) nil
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r nil)
H1:pos_Rl (cons r nil) 0 = a
H2:pos_Rl (cons r nil) (Init.Nat.pred (Rlength (cons r nil))) = b
H3:forall x : R, a < x < b -> f x <= g x

Int_SF (app_Rlist (mid_Rlist nil r) f) (cons r nil) <= Int_SF (app_Rlist (mid_Rlist nil r) g) (cons r nil)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist (cons r0 r1) r) f) (cons r (cons r0 r1)) <= Int_SF (app_Rlist (mid_Rlist (cons r0 r1) r) g) (cons r (cons r0 r1))
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)

Int_SF (app_Rlist (mid_Rlist (cons r0 r1) r) f) (cons r (cons r0 r1)) <= Int_SF (app_Rlist (mid_Rlist (cons r0 r1) r) g) (cons r (cons r0 r1))
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)

f ((r + r0) / 2) * (r0 - r) <= g ((r + r0) / 2) * (r0 - r)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r = r0

f ((r + r0) / 2) * (r0 - r) <= g ((r + r0) / 2) * (r0 - r)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
f ((r + r0) / 2) * (r0 - r) <= g ((r + r0) / 2) * (r0 - r)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0

f ((r + r0) / 2) * (r0 - r) <= g ((r + r0) / 2) * (r0 - r)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0

0 <= r0 - r
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
f ((r + r0) / 2) <= g ((r + r0) / 2)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0

f ((r + r0) / 2) <= g ((r + r0) / 2)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0

a < (r + r0) / 2
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
(r + r0) / 2 < b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0

0 < 2
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 * a < 2 * ((r + r0) / 2)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
(r + r0) / 2 < b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0

2 * a < 2 * ((r + r0) / 2)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
(r + r0) / 2 < b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0

2 * a < 1 * (r + r0)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
(r + r0) / 2 < b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0

r = a
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r = a
2 * a < 1 * (r + r0)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
(r + r0) / 2 < b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r = a

2 * a < 1 * (r + r0)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
(r + r0) / 2 < b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r = a

a < r0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
(r + r0) / 2 < b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r = a
H6:pos_Rl (cons r (cons r0 r1)) 0 <= pos_Rl (cons r (cons r0 r1)) 1

a < r0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
(r + r0) / 2 < b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r = a
H6:r <= r0

a < r0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
(r + r0) / 2 < b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r = a
H6:r <= r0
H7:r < r0

a < r0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r = a
H6:r <= r0
H7:r = r0
a < r0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
(r + r0) / 2 < b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r = a
H6:r <= r0
H7:r = r0

a < r0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
(r + r0) / 2 < b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0

2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
(r + r0) / 2 < b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0

(r + r0) / 2 < b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0

0 < 2
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 * ((r + r0) / 2) < 2 * b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0

2 * ((r + r0) / 2) < 2 * b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0

1 * (r + r0) < 2 * b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0

r0 <= b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b
r + r0 < b + b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0

r0 <= pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1))))
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b
r + r0 < b + b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0

pos_Rl (cons r (cons r0 r1)) 1 <= pos_Rl (cons r (cons (pos_Rl (cons r (cons r0 r1)) 1) r1)) (Init.Nat.pred (Rlength (cons r (cons (pos_Rl (cons r (cons r0 r1)) 1) r1))))
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
pos_Rl (cons r (cons r0 r1)) 1 = r0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b
r + r0 < b + b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:ordered_Rlist (cons r (cons r0 r1)) -> forall i j : nat, (i <= j)%nat -> (j < Rlength (cons r (cons r0 r1)))%nat -> pos_Rl (cons r (cons r0 r1)) i <= pos_Rl (cons r (cons r0 r1)) j
H6:(forall i j : nat, (i <= j)%nat -> (j < Rlength (cons r (cons r0 r1)))%nat -> pos_Rl (cons r (cons r0 r1)) i <= pos_Rl (cons r (cons r0 r1)) j) -> ordered_Rlist (cons r (cons r0 r1))

ordered_Rlist (cons r (cons r0 r1))
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:ordered_Rlist (cons r (cons r0 r1)) -> forall i j : nat, (i <= j)%nat -> (j < Rlength (cons r (cons r0 r1)))%nat -> pos_Rl (cons r (cons r0 r1)) i <= pos_Rl (cons r (cons r0 r1)) j
H6:(forall i j : nat, (i <= j)%nat -> (j < Rlength (cons r (cons r0 r1)))%nat -> pos_Rl (cons r (cons r0 r1)) i <= pos_Rl (cons r (cons r0 r1)) j) -> ordered_Rlist (cons r (cons r0 r1))
(1 <= Init.Nat.pred (Rlength (cons r (cons (pos_Rl (cons r (cons r0 r1)) 1) r1))))%nat
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:ordered_Rlist (cons r (cons r0 r1)) -> forall i j : nat, (i <= j)%nat -> (j < Rlength (cons r (cons r0 r1)))%nat -> pos_Rl (cons r (cons r0 r1)) i <= pos_Rl (cons r (cons r0 r1)) j
H6:(forall i j : nat, (i <= j)%nat -> (j < Rlength (cons r (cons r0 r1)))%nat -> pos_Rl (cons r (cons r0 r1)) i <= pos_Rl (cons r (cons r0 r1)) j) -> ordered_Rlist (cons r (cons r0 r1))
(Init.Nat.pred (Rlength (cons r (cons (pos_Rl (cons r (cons r0 r1)) 1) r1))) < Rlength (cons r (cons r0 r1)))%nat
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
pos_Rl (cons r (cons r0 r1)) 1 = r0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b
r + r0 < b + b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:ordered_Rlist (cons r (cons r0 r1)) -> forall i j : nat, (i <= j)%nat -> (j < Rlength (cons r (cons r0 r1)))%nat -> pos_Rl (cons r (cons r0 r1)) i <= pos_Rl (cons r (cons r0 r1)) j
H6:(forall i j : nat, (i <= j)%nat -> (j < Rlength (cons r (cons r0 r1)))%nat -> pos_Rl (cons r (cons r0 r1)) i <= pos_Rl (cons r (cons r0 r1)) j) -> ordered_Rlist (cons r (cons r0 r1))

(1 <= Init.Nat.pred (Rlength (cons r (cons (pos_Rl (cons r (cons r0 r1)) 1) r1))))%nat
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:ordered_Rlist (cons r (cons r0 r1)) -> forall i j : nat, (i <= j)%nat -> (j < Rlength (cons r (cons r0 r1)))%nat -> pos_Rl (cons r (cons r0 r1)) i <= pos_Rl (cons r (cons r0 r1)) j
H6:(forall i j : nat, (i <= j)%nat -> (j < Rlength (cons r (cons r0 r1)))%nat -> pos_Rl (cons r (cons r0 r1)) i <= pos_Rl (cons r (cons r0 r1)) j) -> ordered_Rlist (cons r (cons r0 r1))
(Init.Nat.pred (Rlength (cons r (cons (pos_Rl (cons r (cons r0 r1)) 1) r1))) < Rlength (cons r (cons r0 r1)))%nat
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
pos_Rl (cons r (cons r0 r1)) 1 = r0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b
r + r0 < b + b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:ordered_Rlist (cons r (cons r0 r1)) -> forall i j : nat, (i <= j)%nat -> (j < Rlength (cons r (cons r0 r1)))%nat -> pos_Rl (cons r (cons r0 r1)) i <= pos_Rl (cons r (cons r0 r1)) j
H6:(forall i j : nat, (i <= j)%nat -> (j < Rlength (cons r (cons r0 r1)))%nat -> pos_Rl (cons r (cons r0 r1)) i <= pos_Rl (cons r (cons r0 r1)) j) -> ordered_Rlist (cons r (cons r0 r1))

(0 <= Rlength r1)%nat
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:ordered_Rlist (cons r (cons r0 r1)) -> forall i j : nat, (i <= j)%nat -> (j < Rlength (cons r (cons r0 r1)))%nat -> pos_Rl (cons r (cons r0 r1)) i <= pos_Rl (cons r (cons r0 r1)) j
H6:(forall i j : nat, (i <= j)%nat -> (j < Rlength (cons r (cons r0 r1)))%nat -> pos_Rl (cons r (cons r0 r1)) i <= pos_Rl (cons r (cons r0 r1)) j) -> ordered_Rlist (cons r (cons r0 r1))
(Init.Nat.pred (Rlength (cons r (cons (pos_Rl (cons r (cons r0 r1)) 1) r1))) < Rlength (cons r (cons r0 r1)))%nat
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
pos_Rl (cons r (cons r0 r1)) 1 = r0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b
r + r0 < b + b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:ordered_Rlist (cons r (cons r0 r1)) -> forall i j : nat, (i <= j)%nat -> (j < Rlength (cons r (cons r0 r1)))%nat -> pos_Rl (cons r (cons r0 r1)) i <= pos_Rl (cons r (cons r0 r1)) j
H6:(forall i j : nat, (i <= j)%nat -> (j < Rlength (cons r (cons r0 r1)))%nat -> pos_Rl (cons r (cons r0 r1)) i <= pos_Rl (cons r (cons r0 r1)) j) -> ordered_Rlist (cons r (cons r0 r1))

(Init.Nat.pred (Rlength (cons r (cons (pos_Rl (cons r (cons r0 r1)) 1) r1))) < Rlength (cons r (cons r0 r1)))%nat
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
pos_Rl (cons r (cons r0 r1)) 1 = r0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b
r + r0 < b + b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0

pos_Rl (cons r (cons r0 r1)) 1 = r0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b
r + r0 < b + b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b

r + r0 < b + b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b

r + r0 <= r + b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b
r + b < b + b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b

r + b < b + b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b

r < b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b

r < r0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b
r0 <= b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b
H6:pos_Rl (cons r (cons r0 r1)) 0 <= pos_Rl (cons r (cons r0 r1)) 1

r < r0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b
r0 <= b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b
H6:r <= r0

r < r0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b
r0 <= b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b
H6:r <= r0
H7:r < r0

r < r0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b
H6:r <= r0
H7:r = r0
r < r0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b
r0 <= b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b
H6:r <= r0
H7:r = r0

r < r0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b
r0 <= b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
H5:r0 <= b

r0 <= b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0
2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
H4:r <> r0

2 <> 0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> pos_Rl (cons r0 r1) 0 = a0 -> pos_Rl (cons r0 r1) (Init.Nat.pred (Rlength (cons r0 r1))) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF (cons r0 r1) f0) (cons r0 r1) <= Int_SF (FF (cons r0 r1) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)

Int_SF (app_Rlist (mid_Rlist r1 r0) f) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g) (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> r0 = a0 -> match Rlength r1 with | 0%nat => r0 | S i' => pos_Rl r1 i' end = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (app_Rlist (mid_Rlist r1 r0) f0) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)

ordered_Rlist (cons r0 r1)
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> r0 = a0 -> match Rlength r1 with | 0%nat => r0 | S i' => pos_Rl r1 i' end = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (app_Rlist (mid_Rlist r1 r0) f0) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
r0 = r0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> r0 = a0 -> match Rlength r1 with | 0%nat => r0 | S i' => pos_Rl r1 i' end = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (app_Rlist (mid_Rlist r1 r0) f0) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
match Rlength r1 with | 0%nat => r0 | S i' => pos_Rl r1 i' end = b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> r0 = a0 -> match Rlength r1 with | 0%nat => r0 | S i' => pos_Rl r1 i' end = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (app_Rlist (mid_Rlist r1 r0) f0) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
forall x : R, r0 < x < b -> f x <= g x
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> r0 = a0 -> match Rlength r1 with | 0%nat => r0 | S i' => pos_Rl r1 i' end = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (app_Rlist (mid_Rlist r1 r0) f0) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)

r0 = r0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> r0 = a0 -> match Rlength r1 with | 0%nat => r0 | S i' => pos_Rl r1 i' end = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (app_Rlist (mid_Rlist r1 r0) f0) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
match Rlength r1 with | 0%nat => r0 | S i' => pos_Rl r1 i' end = b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> r0 = a0 -> match Rlength r1 with | 0%nat => r0 | S i' => pos_Rl r1 i' end = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (app_Rlist (mid_Rlist r1 r0) f0) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
forall x : R, r0 < x < b -> f x <= g x
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> r0 = a0 -> match Rlength r1 with | 0%nat => r0 | S i' => pos_Rl r1 i' end = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (app_Rlist (mid_Rlist r1 r0) f0) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)

match Rlength r1 with | 0%nat => r0 | S i' => pos_Rl r1 i' end = b
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> r0 = a0 -> match Rlength r1 with | 0%nat => r0 | S i' => pos_Rl r1 i' end = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (app_Rlist (mid_Rlist r1 r0) f0) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
forall x : R, r0 < x < b -> f x <= g x
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> r0 = a0 -> match Rlength r1 with | 0%nat => r0 | S i' => pos_Rl r1 i' end = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (app_Rlist (mid_Rlist r1 r0) f0) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x : R, a < x < b -> f x <= g x
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x : R, a0 < x < b0 -> f0 x <= g0 x) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)

forall x : R, r0 < x < b -> f x <= g x
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> r0 = a0 -> match Rlength r1 with | 0%nat => r0 | S i' => pos_Rl r1 i' end = b0 -> (forall x0 : R, a0 < x0 < b0 -> f0 x0 <= g0 x0) -> Int_SF (app_Rlist (mid_Rlist r1 r0) f0) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x0 : R, a < x0 < b -> f x0 <= g x0
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x0 : R, a0 < x0 < b0 -> f0 x0 <= g0 x0) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
x:R
H4:r0 < x < b
H5:r0 < x
H6:x < b

a < x
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> r0 = a0 -> match Rlength r1 with | 0%nat => r0 | S i' => pos_Rl r1 i' end = b0 -> (forall x0 : R, a0 < x0 < b0 -> f0 x0 <= g0 x0) -> Int_SF (app_Rlist (mid_Rlist r1 r0) f0) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x0 : R, a < x0 < b -> f x0 <= g x0
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x0 : R, a0 < x0 < b0 -> f0 x0 <= g0 x0) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
x:R
H4:r0 < x < b
H5:r0 < x
H6:x < b

a <= r0
l:Rlist
r, r0:R
r1:Rlist
H:forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist (cons r0 r1) -> r0 = a0 -> match Rlength r1 with | 0%nat => r0 | S i' => pos_Rl r1 i' end = b0 -> (forall x0 : R, a0 < x0 < b0 -> f0 x0 <= g0 x0) -> Int_SF (app_Rlist (mid_Rlist r1 r0) f0) (cons r0 r1) <= Int_SF (app_Rlist (mid_Rlist r1 r0) g0) (cons r0 r1)
a, b:R
f, g:R -> R
H0:ordered_Rlist (cons r (cons r0 r1))
H1:pos_Rl (cons r (cons r0 r1)) 0 = a
H2:pos_Rl (cons r (cons r0 r1)) (Init.Nat.pred (Rlength (cons r (cons r0 r1)))) = b
H3:forall x0 : R, a < x0 < b -> f x0 <= g x0
Hrecr0:(forall (a0 b0 : R) (f0 g0 : R -> R), ordered_Rlist r1 -> pos_Rl r1 0 = a0 -> pos_Rl r1 (Init.Nat.pred (Rlength r1)) = b0 -> (forall x0 : R, a0 < x0 < b0 -> f0 x0 <= g0 x0) -> Int_SF (FF r1 f0) r1 <= Int_SF (FF r1 g0) r1) -> ordered_Rlist (cons r r1) -> pos_Rl (cons r r1) 0 = a -> pos_Rl (cons r r1) (Init.Nat.pred (Rlength (cons r r1))) = b -> Int_SF (app_Rlist (mid_Rlist r1 r) f) (cons r r1) <= Int_SF (app_Rlist (mid_Rlist r1 r) g) (cons r r1)
x:R
H4:r0 < x < b
H5:r0 < x
H6:x < b

pos_Rl (cons r (cons r0 r1)) 0 <= r0
simpl; apply (H0 0%nat); simpl; apply lt_O_Sn. Qed.

forall (a b : R) (f g : StepFun a b) (l : Rlist), a <= b -> is_subdivision f a b l -> is_subdivision g a b l -> (forall x : R, a < x < b -> f x <= g x) -> RiemannInt_SF f <= RiemannInt_SF g

forall (a b : R) (f g : StepFun a b) (l : Rlist), a <= b -> is_subdivision f a b l -> is_subdivision g a b l -> (forall x : R, a < x < b -> f x <= g x) -> RiemannInt_SF f <= RiemannInt_SF g
a, b:R
f, g:StepFun a b
l:Rlist
X:is_subdivision f a b l
X0:is_subdivision g a b l
H0:forall x : R, a < x < b -> f x <= g x
H:a <= b

Int_SF (subdivision_val f) (subdivision f) <= Int_SF (subdivision_val g) (subdivision g)
a, b:R
f, g:StepFun a b
l:Rlist
X:is_subdivision f a b l
X0:is_subdivision g a b l
H0:forall x : R, a < x < b -> f x <= g x
H:a <= b

Int_SF (FF l f) l <= Int_SF (subdivision_val g) (subdivision g)
a, b:R
f, g:StepFun a b
l:Rlist
X:is_subdivision f a b l
X0:is_subdivision g a b l
H0:forall x : R, a < x < b -> f x <= g x
H:a <= b
Int_SF (FF l f) l = Int_SF (subdivision_val f) (subdivision f)
a, b:R
f, g:StepFun a b
l:Rlist
X:is_subdivision f a b l
X0:is_subdivision g a b l
H0:forall x : R, a < x < b -> f x <= g x
H:a <= b

Int_SF (FF l f) l <= Int_SF (FF l g) l
a, b:R
f, g:StepFun a b
l:Rlist
X:is_subdivision f a b l
X0:is_subdivision g a b l
H0:forall x : R, a < x < b -> f x <= g x
H:a <= b
Int_SF (FF l g) l = Int_SF (subdivision_val g) (subdivision g)
a, b:R
f, g:StepFun a b
l:Rlist
X:is_subdivision f a b l
X0:is_subdivision g a b l
H0:forall x : R, a < x < b -> f x <= g x
H:a <= b
Int_SF (FF l f) l = Int_SF (subdivision_val f) (subdivision f)
a, b:R
f, g:StepFun a b
l:Rlist
X:is_subdivision f a b l
X0:is_subdivision g a b l
H0:forall x : R, a < x < b -> f x <= g x
H:a <= b

Int_SF (FF l g) l = Int_SF (subdivision_val g) (subdivision g)
a, b:R
f, g:StepFun a b
l:Rlist
X:is_subdivision f a b l
X0:is_subdivision g a b l
H0:forall x : R, a < x < b -> f x <= g x
H:a <= b
Int_SF (FF l f) l = Int_SF (subdivision_val f) (subdivision f)
a, b:R
f, g:StepFun a b
l:Rlist
X:is_subdivision f a b l
X0:is_subdivision g a b l
H0:forall x : R, a < x < b -> f x <= g x
H:a <= b

Int_SF (FF l f) l = Int_SF (subdivision_val f) (subdivision f)
apply StepFun_P17 with (fe f) a b; [ apply StepFun_P21; assumption | apply StepFun_P1 ]. Qed.

forall (a b : R) (f g : StepFun a b), a <= b -> (forall x : R, a < x < b -> f x <= g x) -> RiemannInt_SF f <= RiemannInt_SF g

forall (a b : R) (f g : StepFun a b), a <= b -> (forall x : R, a < x < b -> f x <= g x) -> RiemannInt_SF f <= RiemannInt_SF g
a, b:R
f, g:StepFun a b
H:a <= b
H0:forall x : R, a < x < b -> f x <= g x

is_subdivision f a b ?l
a, b:R
f, g:StepFun a b
H:a <= b
H0:forall x : R, a < x < b -> f x <= g x
is_subdivision g a b ?l
a, b:R
f, g:StepFun a b
H:a <= b
H0:forall x : R, a < x < b -> f x <= g x

is_subdivision g a b (cons_ORlist (subdivision ?f) (subdivision f))
eapply StepFun_P23; apply StepFun_P29. Qed.

forall (l : Rlist) (a b : R) (f : R -> R), ordered_Rlist l -> pos_Rl l 0 = a -> pos_Rl l (Init.Nat.pred (Rlength l)) = b -> {g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq g (co_interval (pos_Rl l i) (pos_Rl l (S i))) (f (pos_Rl l i)))}

forall (l : Rlist) (a b : R) (f : R -> R), ordered_Rlist l -> pos_Rl l 0 = a -> pos_Rl l (Init.Nat.pred (Rlength l)) = b -> {g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq g (co_interval (pos_Rl l i) (pos_Rl l (S i))) (f (pos_Rl l i)))}
b:R
f:R -> R

forall a : R, ordered_Rlist nil -> pos_Rl nil 0 = a -> pos_Rl nil (Init.Nat.pred (Rlength nil)) = b -> {g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength nil))%nat -> constant_D_eq g (co_interval (pos_Rl nil i) (pos_Rl nil (S i))) (f (pos_Rl nil i)))}
r:R
l:Rlist
b:R
f:R -> R
IHl:forall a : R, ordered_Rlist l -> pos_Rl l 0 = a -> pos_Rl l (Init.Nat.pred (Rlength l)) = b -> {g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq g (co_interval (pos_Rl l i) (pos_Rl l (S i))) (f (pos_Rl l i)))}
forall a : R, ordered_Rlist (cons r l) -> pos_Rl (cons r l) 0 = a -> pos_Rl (cons r l) (Init.Nat.pred (Rlength (cons r l))) = b -> {g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r l) i) (pos_Rl (cons r l) (S i))) (f (pos_Rl (cons r l) i)))}
b:R
f:R -> R
a:R
H:ordered_Rlist nil
H0:0 = a
H1:0 = b

{| fe := fct_cte (f b); pre := StepFun_P4 a b (f b) |} b = f b
b:R
f:R -> R
a:R
H:ordered_Rlist nil
H0:0 = a
H1:0 = b
forall i : nat, (i < Init.Nat.pred (Rlength nil))%nat -> constant_D_eq {| fe := fct_cte (f b); pre := StepFun_P4 a b (f b) |} (co_interval (pos_Rl nil i) (pos_Rl nil (S i))) (f (pos_Rl nil i))
r:R
l:Rlist
b:R
f:R -> R
IHl:forall a : R, ordered_Rlist l -> pos_Rl l 0 = a -> pos_Rl l (Init.Nat.pred (Rlength l)) = b -> {g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq g (co_interval (pos_Rl l i) (pos_Rl l (S i))) (f (pos_Rl l i)))}
forall a : R, ordered_Rlist (cons r l) -> pos_Rl (cons r l) 0 = a -> pos_Rl (cons r l) (Init.Nat.pred (Rlength (cons r l))) = b -> {g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r l) i) (pos_Rl (cons r l) (S i))) (f (pos_Rl (cons r l) i)))}
b:R
f:R -> R
a:R
H:ordered_Rlist nil
H0:0 = a
H1:0 = b

forall i : nat, (i < Init.Nat.pred (Rlength nil))%nat -> constant_D_eq {| fe := fct_cte (f b); pre := StepFun_P4 a b (f b) |} (co_interval (pos_Rl nil i) (pos_Rl nil (S i))) (f (pos_Rl nil i))
r:R
l:Rlist
b:R
f:R -> R
IHl:forall a : R, ordered_Rlist l -> pos_Rl l 0 = a -> pos_Rl l (Init.Nat.pred (Rlength l)) = b -> {g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq g (co_interval (pos_Rl l i) (pos_Rl l (S i))) (f (pos_Rl l i)))}
forall a : R, ordered_Rlist (cons r l) -> pos_Rl (cons r l) 0 = a -> pos_Rl (cons r l) (Init.Nat.pred (Rlength (cons r l))) = b -> {g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r l) i) (pos_Rl (cons r l) (S i))) (f (pos_Rl (cons r l) i)))}
r:R
l:Rlist
b:R
f:R -> R
IHl:forall a : R, ordered_Rlist l -> pos_Rl l 0 = a -> pos_Rl l (Init.Nat.pred (Rlength l)) = b -> {g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength l))%nat -> constant_D_eq g (co_interval (pos_Rl l i) (pos_Rl l (S i))) (f (pos_Rl l i)))}

forall a : R, ordered_Rlist (cons r l) -> pos_Rl (cons r l) 0 = a -> pos_Rl (cons r l) (Init.Nat.pred (Rlength (cons r l))) = b -> {g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r l) i) (pos_Rl (cons r l) (S i))) (f (pos_Rl (cons r l) i)))}
r, b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist nil -> pos_Rl nil 0 = a0 -> pos_Rl nil (Init.Nat.pred (Rlength nil)) = b -> {g : StepFun a0 b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength nil))%nat -> constant_D_eq g (co_interval (pos_Rl nil i) (pos_Rl nil (S i))) (f (pos_Rl nil i)))}
a:R
H:ordered_Rlist (cons r nil)
H0:pos_Rl (cons r nil) 0 = a
H1:pos_Rl (cons r nil) (Init.Nat.pred (Rlength (cons r nil))) = b

{g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (f (pos_Rl (cons r nil) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g : StepFun a0 b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
{g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist nil -> pos_Rl nil 0 = a0 -> pos_Rl nil (Init.Nat.pred (Rlength nil)) = b -> {g : StepFun a0 b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength nil))%nat -> constant_D_eq g (co_interval (pos_Rl nil i) (pos_Rl nil (S i))) (f (pos_Rl nil i)))}
a:R
H:ordered_Rlist (cons r nil)
H0:r = a
H1:r = b

{| fe := fct_cte (f b); pre := StepFun_P4 a b (f b) |} b = f b
r, b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist nil -> pos_Rl nil 0 = a0 -> pos_Rl nil (Init.Nat.pred (Rlength nil)) = b -> {g : StepFun a0 b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength nil))%nat -> constant_D_eq g (co_interval (pos_Rl nil i) (pos_Rl nil (S i))) (f (pos_Rl nil i)))}
a:R
H:ordered_Rlist (cons r nil)
H0:r = a
H1:r = b
forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq {| fe := fct_cte (f b); pre := StepFun_P4 a b (f b) |} (co_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (f (pos_Rl (cons r nil) i))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g : StepFun a0 b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
{g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist nil -> pos_Rl nil 0 = a0 -> pos_Rl nil (Init.Nat.pred (Rlength nil)) = b -> {g : StepFun a0 b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength nil))%nat -> constant_D_eq g (co_interval (pos_Rl nil i) (pos_Rl nil (S i))) (f (pos_Rl nil i)))}
a:R
H:ordered_Rlist (cons r nil)
H0:r = a
H1:r = b

forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq {| fe := fct_cte (f b); pre := StepFun_P4 a b (f b) |} (co_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (f (pos_Rl (cons r nil) i))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g : StepFun a0 b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
{g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g : StepFun a0 b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b

{g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g : StepFun a0 b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b

ordered_Rlist (cons r1 l)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g : StepFun a0 b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
{g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g : StepFun a0 b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)

{g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g : StepFun a0 b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)

pos_Rl (cons r1 l) 0 = r1
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g : StepFun a0 b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
{g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g : StepFun a0 b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1

{g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g : StepFun a0 b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1

pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g : StepFun a0 b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
{g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g : StepFun a0 b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b

{g : StepFun a b | g b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))

{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R

{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R

r1 <= b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b

{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b

IsStepFun g' a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)

ordered_Rlist (cons a lg)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
pos_Rl (cons a lg) 0 = Rmin a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
pos_Rl (cons a lg) (Init.Nat.pred (Rlength (cons a lg))) = Rmax a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
Rlength (cons a lg) = S (Rlength (cons (f a) lg2))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons a lg)))%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H9:(0 < Rlength lg)%nat

pos_Rl (cons a lg) 0 <= pos_Rl (cons a lg) 1
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
Hreci:(i < Rlength lg)%nat -> pos_Rl (cons a lg) i <= pos_Rl (cons a lg) (S i)
pos_Rl (cons a lg) (S i) <= pos_Rl (cons a lg) (S (S i))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
pos_Rl (cons a lg) 0 = Rmin a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
pos_Rl (cons a lg) (Init.Nat.pred (Rlength (cons a lg))) = Rmax a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
Rlength (cons a lg) = S (Rlength (cons (f a) lg2))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons a lg)))%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H9:(0 < Rlength lg)%nat

a <= r1
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H9:(0 < Rlength lg)%nat
r1 = Rmin r1 b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
Hreci:(i < Rlength lg)%nat -> pos_Rl (cons a lg) i <= pos_Rl (cons a lg) (S i)
pos_Rl (cons a lg) (S i) <= pos_Rl (cons a lg) (S (S i))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
pos_Rl (cons a lg) 0 = Rmin a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
pos_Rl (cons a lg) (Init.Nat.pred (Rlength (cons a lg))) = Rmax a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
Rlength (cons a lg) = S (Rlength (cons (f a) lg2))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons a lg)))%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H9:(0 < Rlength lg)%nat

r1 = Rmin r1 b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
Hreci:(i < Rlength lg)%nat -> pos_Rl (cons a lg) i <= pos_Rl (cons a lg) (S i)
pos_Rl (cons a lg) (S i) <= pos_Rl (cons a lg) (S (S i))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
pos_Rl (cons a lg) 0 = Rmin a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
pos_Rl (cons a lg) (Init.Nat.pred (Rlength (cons a lg))) = Rmax a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
Rlength (cons a lg) = S (Rlength (cons (f a) lg2))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons a lg)))%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
Hreci:(i < Rlength lg)%nat -> pos_Rl (cons a lg) i <= pos_Rl (cons a lg) (S i)

pos_Rl (cons a lg) (S i) <= pos_Rl (cons a lg) (S (S i))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
pos_Rl (cons a lg) 0 = Rmin a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
pos_Rl (cons a lg) (Init.Nat.pred (Rlength (cons a lg))) = Rmax a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
Rlength (cons a lg) = S (Rlength (cons (f a) lg2))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons a lg)))%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
Hreci:(i < Rlength lg)%nat -> pos_Rl (cons a lg) i <= pos_Rl (cons a lg) (S i)

(S i < S (Init.Nat.pred (Rlength lg)))%nat
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
pos_Rl (cons a lg) 0 = Rmin a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
pos_Rl (cons a lg) (Init.Nat.pred (Rlength (cons a lg))) = Rmax a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
Rlength (cons a lg) = S (Rlength (cons (f a) lg2))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons a lg)))%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
Hreci:(i < Rlength lg)%nat -> pos_Rl (cons a lg) i <= pos_Rl (cons a lg) (S i)

(S i < Rlength lg)%nat
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
Hreci:(i < Rlength lg)%nat -> pos_Rl (cons a lg) i <= pos_Rl (cons a lg) (S i)
Rlength lg = S (Init.Nat.pred (Rlength lg))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
pos_Rl (cons a lg) 0 = Rmin a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
pos_Rl (cons a lg) (Init.Nat.pred (Rlength (cons a lg))) = Rmax a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
Rlength (cons a lg) = S (Rlength (cons (f a) lg2))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons a lg)))%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
Hreci:(i < Rlength lg)%nat -> pos_Rl (cons a lg) i <= pos_Rl (cons a lg) (S i)

Rlength lg = S (Init.Nat.pred (Rlength lg))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
pos_Rl (cons a lg) 0 = Rmin a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
pos_Rl (cons a lg) (Init.Nat.pred (Rlength (cons a lg))) = Rmax a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
Rlength (cons a lg) = S (Rlength (cons (f a) lg2))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons a lg)))%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)

pos_Rl (cons a lg) 0 = Rmin a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
pos_Rl (cons a lg) (Init.Nat.pred (Rlength (cons a lg))) = Rmax a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
Rlength (cons a lg) = S (Rlength (cons (f a) lg2))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons a lg)))%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)

a <= b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H14:a <= b
a = Rmin a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
pos_Rl (cons a lg) (Init.Nat.pred (Rlength (cons a lg))) = Rmax a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
Rlength (cons a lg) = S (Rlength (cons (f a) lg2))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons a lg)))%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H14:a <= b

a = Rmin a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
pos_Rl (cons a lg) (Init.Nat.pred (Rlength (cons a lg))) = Rmax a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
Rlength (cons a lg) = S (Rlength (cons (f a) lg2))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons a lg)))%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)

pos_Rl (cons a lg) (Init.Nat.pred (Rlength (cons a lg))) = Rmax a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
Rlength (cons a lg) = S (Rlength (cons (f a) lg2))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons a lg)))%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)

a <= b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H14:a <= b
pos_Rl (cons a lg) (Init.Nat.pred (Rlength (cons a lg))) = Rmax a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
Rlength (cons a lg) = S (Rlength (cons (f a) lg2))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons a lg)))%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H14:a <= b

pos_Rl (cons a lg) (Init.Nat.pred (Rlength (cons a lg))) = Rmax a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
Rlength (cons a lg) = S (Rlength (cons (f a) lg2))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons a lg)))%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H14:a <= b

pos_Rl (cons a lg) (Init.Nat.pred (Rlength (cons a lg))) = Rmax r1 b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H14:a <= b
Rmax r1 b = Rmax a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
Rlength (cons a lg) = S (Rlength (cons (f a) lg2))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons a lg)))%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg2:Rlist
H10:ordered_Rlist nil
H12:pos_Rl nil 0 = Rmin r1 b
H11:pos_Rl nil (Init.Nat.pred (Rlength nil)) = Rmax r1 b
H13:Rlength nil = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength nil))%nat -> constant_D_eq g (open_interval (pos_Rl nil i) (pos_Rl nil (S i))) (pos_Rl lg2 i)
H14:a <= b

pos_Rl (cons a nil) (Init.Nat.pred (Rlength (cons a nil))) = pos_Rl nil (Init.Nat.pred (Rlength nil))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
r0:R
lg, lg2:Rlist
H10:ordered_Rlist (cons r0 lg)
H12:pos_Rl (cons r0 lg) 0 = Rmin r1 b
H11:pos_Rl (cons r0 lg) (Init.Nat.pred (Rlength (cons r0 lg))) = Rmax r1 b
H13:Rlength (cons r0 lg) = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r0 lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons r0 lg) i) (pos_Rl (cons r0 lg) (S i))) (pos_Rl lg2 i)
H14:a <= b
Hreclg:ordered_Rlist lg -> pos_Rl lg 0 = Rmin r1 b -> pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b -> Rlength lg = S (Rlength lg2) -> (forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)) -> pos_Rl (cons a lg) (Init.Nat.pred (Rlength (cons a lg))) = pos_Rl lg (Init.Nat.pred (Rlength lg))
pos_Rl (cons a (cons r0 lg)) (Init.Nat.pred (Rlength (cons a (cons r0 lg)))) = pos_Rl (cons r0 lg) (Init.Nat.pred (Rlength (cons r0 lg)))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H14:a <= b
Rmax r1 b = Rmax a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
Rlength (cons a lg) = S (Rlength (cons (f a) lg2))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons a lg)))%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
r0:R
lg, lg2:Rlist
H10:ordered_Rlist (cons r0 lg)
H12:pos_Rl (cons r0 lg) 0 = Rmin r1 b
H11:pos_Rl (cons r0 lg) (Init.Nat.pred (Rlength (cons r0 lg))) = Rmax r1 b
H13:Rlength (cons r0 lg) = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r0 lg)))%nat -> constant_D_eq g (open_interval (pos_Rl (cons r0 lg) i) (pos_Rl (cons r0 lg) (S i))) (pos_Rl lg2 i)
H14:a <= b
Hreclg:ordered_Rlist lg -> pos_Rl lg 0 = Rmin r1 b -> pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b -> Rlength lg = S (Rlength lg2) -> (forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)) -> pos_Rl (cons a lg) (Init.Nat.pred (Rlength (cons a lg))) = pos_Rl lg (Init.Nat.pred (Rlength lg))

pos_Rl (cons a (cons r0 lg)) (Init.Nat.pred (Rlength (cons a (cons r0 lg)))) = pos_Rl (cons r0 lg) (Init.Nat.pred (Rlength (cons r0 lg)))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H14:a <= b
Rmax r1 b = Rmax a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
Rlength (cons a lg) = S (Rlength (cons (f a) lg2))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons a lg)))%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H14:a <= b

Rmax r1 b = Rmax a b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
Rlength (cons a lg) = S (Rlength (cons (f a) lg2))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons a lg)))%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)

Rlength (cons a lg) = S (Rlength (cons (f a) lg2))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons a lg)))%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)

forall i : nat, (i < Init.Nat.pred (Rlength (cons a lg)))%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H9:(0 < Rlength lg)%nat

constant_D_eq g' (open_interval (pos_Rl (cons a lg) 0) (pos_Rl (cons a lg) 1)) (pos_Rl (cons (f a) lg2) 0)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
Hreci:(i < Rlength lg)%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
constant_D_eq g' (open_interval (pos_Rl (cons a lg) (S i)) (pos_Rl (cons a lg) (S (S i)))) (pos_Rl (cons (f a) lg2) (S i))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H9:(0 < Rlength lg)%nat
x:R
H14:a < x < pos_Rl lg 0

Rmin r1 b = r1
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H9:(0 < Rlength lg)%nat
x:R
H14:a < x < pos_Rl lg 0
H16:Rmin r1 b = r1
g' x = f a
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
Hreci:(i < Rlength lg)%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
constant_D_eq g' (open_interval (pos_Rl (cons a lg) (S i)) (pos_Rl (cons a lg) (S (S i)))) (pos_Rl (cons (f a) lg2) (S i))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H9:(0 < Rlength lg)%nat
x:R
H14:a < x < pos_Rl lg 0
H16:Rmin r1 b = r1

g' x = f a
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
Hreci:(i < Rlength lg)%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
constant_D_eq g' (open_interval (pos_Rl (cons a lg) (S i)) (pos_Rl (cons a lg) (S (S i)))) (pos_Rl (cons (f a) lg2) (S i))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = r1
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H9:(0 < Rlength lg)%nat
x:R
H16:Rmin r1 b = r1
H14:x < r1
r3:r1 <= x

g x = f a
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = r1
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H9:(0 < Rlength lg)%nat
x:R
H16:Rmin r1 b = r1
H14:x < r1
r3:~ r1 <= x
f a = f a
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
Hreci:(i < Rlength lg)%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
constant_D_eq g' (open_interval (pos_Rl (cons a lg) (S i)) (pos_Rl (cons a lg) (S (S i)))) (pos_Rl (cons (f a) lg2) (S i))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = r1
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i : nat, (i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H9:(0 < Rlength lg)%nat
x:R
H16:Rmin r1 b = r1
H14:x < r1
r3:~ r1 <= x

f a = f a
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
Hreci:(i < Rlength lg)%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)
constant_D_eq g' (open_interval (pos_Rl (cons a lg) (S i)) (pos_Rl (cons a lg) (S (S i)))) (pos_Rl (cons (f a) lg2) (S i))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
Hreci:(i < Rlength lg)%nat -> constant_D_eq g' (open_interval (pos_Rl (cons a lg) i) (pos_Rl (cons a lg) (S i))) (pos_Rl (cons (f a) lg2) i)

constant_D_eq g' (open_interval (pos_Rl (cons a lg) (S i)) (pos_Rl (cons a lg) (S (S i)))) (pos_Rl (cons (f a) lg2) (S i))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)

(i < Init.Nat.pred (Rlength lg))%nat
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
constant_D_eq g' (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)

(S i < S (Init.Nat.pred (Rlength lg)))%nat
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
constant_D_eq g' (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)

(S i < Rlength lg)%nat
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
Rlength lg = S (Init.Nat.pred (Rlength lg))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
constant_D_eq g' (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)

Rlength lg = S (Init.Nat.pred (Rlength lg))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
constant_D_eq g' (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat

constant_D_eq g' (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
H18:forall x0 : R, pos_Rl lg i < x0 < pos_Rl lg (S i) -> g x0 = pos_Rl lg2 i
x:R
H14:pos_Rl lg i < x < pos_Rl lg (S i)
H19:g x = pos_Rl lg2 i
r0:r1 <= x

g x = g x
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
H18:forall x0 : R, pos_Rl lg i < x0 < pos_Rl lg (S i) -> g x0 = pos_Rl lg2 i
x:R
H14:pos_Rl lg i < x < pos_Rl lg (S i)
H19:g x = pos_Rl lg2 i
r1 <= x
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
H18:forall x0 : R, pos_Rl lg i < x0 < pos_Rl lg (S i) -> g x0 = pos_Rl lg2 i
x:R
H14:pos_Rl lg i < x < pos_Rl lg (S i)
H19:g x = pos_Rl lg2 i

r1 <= x
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
H18:forall x0 : R, pos_Rl lg i < x0 < pos_Rl lg (S i) -> g x0 = pos_Rl lg2 i
x:R
H14:pos_Rl lg i < x < pos_Rl lg (S i)
H19:g x = pos_Rl lg2 i

Rmin r1 b <= x
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
H18:forall x0 : R, pos_Rl lg i < x0 < pos_Rl lg (S i) -> g x0 = pos_Rl lg2 i
x:R
H14:pos_Rl lg i < x < pos_Rl lg (S i)
H19:g x = pos_Rl lg2 i
Rmin r1 b = r1
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
H18:forall x0 : R, pos_Rl lg i < x0 < pos_Rl lg (S i) -> g x0 = pos_Rl lg2 i
x:R
H19:g x = pos_Rl lg2 i
H14:pos_Rl lg i < x

pos_Rl lg 0 <= pos_Rl lg i
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
H18:forall x0 : R, pos_Rl lg i < x0 < pos_Rl lg (S i) -> g x0 = pos_Rl lg2 i
x:R
H14:pos_Rl lg i < x < pos_Rl lg (S i)
H19:g x = pos_Rl lg2 i
Rmin r1 b = r1
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
H18:forall x0 : R, pos_Rl lg i < x0 < pos_Rl lg (S i) -> g x0 = pos_Rl lg2 i
x:R
H19:g x = pos_Rl lg2 i
H14:pos_Rl lg i < x

ordered_Rlist lg
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
H18:forall x0 : R, pos_Rl lg i < x0 < pos_Rl lg (S i) -> g x0 = pos_Rl lg2 i
x:R
H19:g x = pos_Rl lg2 i
H14:pos_Rl lg i < x
In (pos_Rl lg i) lg
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
H18:forall x0 : R, pos_Rl lg i < x0 < pos_Rl lg (S i) -> g x0 = pos_Rl lg2 i
x:R
H14:pos_Rl lg i < x < pos_Rl lg (S i)
H19:g x = pos_Rl lg2 i
Rmin r1 b = r1
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
H18:forall x0 : R, pos_Rl lg i < x0 < pos_Rl lg (S i) -> g x0 = pos_Rl lg2 i
x:R
H19:g x = pos_Rl lg2 i
H14:pos_Rl lg i < x

In (pos_Rl lg i) lg
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
H18:forall x0 : R, pos_Rl lg i < x0 < pos_Rl lg (S i) -> g x0 = pos_Rl lg2 i
x:R
H14:pos_Rl lg i < x < pos_Rl lg (S i)
H19:g x = pos_Rl lg2 i
Rmin r1 b = r1
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
H18:forall x0 : R, pos_Rl lg i < x0 < pos_Rl lg (S i) -> g x0 = pos_Rl lg2 i
x:R
H19:g x = pos_Rl lg2 i
H14:pos_Rl lg i < x
H20:In (pos_Rl lg i) lg -> exists i0 : nat, pos_Rl lg i = pos_Rl lg i0 /\ (i0 < Rlength lg)%nat
H21:(exists i0 : nat, pos_Rl lg i = pos_Rl lg i0 /\ (i0 < Rlength lg)%nat) -> In (pos_Rl lg i) lg

pos_Rl lg i = pos_Rl lg i
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
H18:forall x0 : R, pos_Rl lg i < x0 < pos_Rl lg (S i) -> g x0 = pos_Rl lg2 i
x:R
H19:g x = pos_Rl lg2 i
H14:pos_Rl lg i < x
H20:In (pos_Rl lg i) lg -> exists i0 : nat, pos_Rl lg i = pos_Rl lg i0 /\ (i0 < Rlength lg)%nat
H21:(exists i0 : nat, pos_Rl lg i = pos_Rl lg i0 /\ (i0 < Rlength lg)%nat) -> In (pos_Rl lg i) lg
(i < Rlength lg)%nat
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
H18:forall x0 : R, pos_Rl lg i < x0 < pos_Rl lg (S i) -> g x0 = pos_Rl lg2 i
x:R
H14:pos_Rl lg i < x < pos_Rl lg (S i)
H19:g x = pos_Rl lg2 i
Rmin r1 b = r1
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
H18:forall x0 : R, pos_Rl lg i < x0 < pos_Rl lg (S i) -> g x0 = pos_Rl lg2 i
x:R
H19:g x = pos_Rl lg2 i
H14:pos_Rl lg i < x
H20:In (pos_Rl lg i) lg -> exists i0 : nat, pos_Rl lg i = pos_Rl lg i0 /\ (i0 < Rlength lg)%nat
H21:(exists i0 : nat, pos_Rl lg i = pos_Rl lg i0 /\ (i0 < Rlength lg)%nat) -> In (pos_Rl lg i) lg

(i < Rlength lg)%nat
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
H18:forall x0 : R, pos_Rl lg i < x0 < pos_Rl lg (S i) -> g x0 = pos_Rl lg2 i
x:R
H14:pos_Rl lg i < x < pos_Rl lg (S i)
H19:g x = pos_Rl lg2 i
Rmin r1 b = r1
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
H18:forall x0 : R, pos_Rl lg i < x0 < pos_Rl lg (S i) -> g x0 = pos_Rl lg2 i
x:R
H19:g x = pos_Rl lg2 i
H14:pos_Rl lg i < x
H20:In (pos_Rl lg i) lg -> exists i0 : nat, pos_Rl lg i = pos_Rl lg i0 /\ (i0 < Rlength lg)%nat
H21:(exists i0 : nat, pos_Rl lg i = pos_Rl lg i0 /\ (i0 < Rlength lg)%nat) -> In (pos_Rl lg i) lg

(Init.Nat.pred (Rlength lg) < Rlength lg)%nat
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
H18:forall x0 : R, pos_Rl lg i < x0 < pos_Rl lg (S i) -> g x0 = pos_Rl lg2 i
x:R
H14:pos_Rl lg i < x < pos_Rl lg (S i)
H19:g x = pos_Rl lg2 i
Rmin r1 b = r1
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:{l0 : Rlist & is_subdivision g r1 b l0}
lg, lg2:Rlist
H10:ordered_Rlist lg
H12:pos_Rl lg 0 = Rmin r1 b
H11:pos_Rl lg (Init.Nat.pred (Rlength lg)) = Rmax r1 b
H13:Rlength lg = S (Rlength lg2)
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i0) (pos_Rl lg (S i0))) (pos_Rl lg2 i0)
i:nat
H9:(S i < Rlength lg)%nat
H16:(i < Init.Nat.pred (Rlength lg))%nat -> constant_D_eq g (open_interval (pos_Rl lg i) (pos_Rl lg (S i))) (pos_Rl lg2 i)
H17:(i < Init.Nat.pred (Rlength lg))%nat
H18:forall x0 : R, pos_Rl lg i < x0 < pos_Rl lg (S i) -> g x0 = pos_Rl lg2 i
x:R
H14:pos_Rl lg i < x < pos_Rl lg (S i)
H19:g x = pos_Rl lg2 i

Rmin r1 b = r1
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b

{g0 : StepFun a b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i)))}
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b

{| fe := g'; pre := H8 |} b = f b
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq {| fe := g'; pre := H8 |} (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b

forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 l))))%nat -> constant_D_eq {| fe := g'; pre := H8 |} (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
H9:(0 < S (Rlength l))%nat

constant_D_eq {| fe := g'; pre := H8 |} (co_interval (pos_Rl (cons r (cons r1 l)) 0) (pos_Rl (cons r (cons r1 l)) 1)) (f (pos_Rl (cons r (cons r1 l)) 0))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
i:nat
H9:(S i < S (Rlength l))%nat
Hreci:(i < S (Rlength l))%nat -> constant_D_eq {| fe := g'; pre := H8 |} (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i))
constant_D_eq {| fe := g'; pre := H8 |} (co_interval (pos_Rl (cons r (cons r1 l)) (S i)) (pos_Rl (cons r (cons r1 l)) (S (S i)))) (f (pos_Rl (cons r (cons r1 l)) (S i)))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:r = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
H9:(0 < S (Rlength l))%nat
x:R
H10:r <= x
H11:x < r1
r3:r1 <= x

g x = f a
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:r = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
H9:(0 < S (Rlength l))%nat
x:R
H10:r <= x
H11:x < r1
r3:~ r1 <= x
f a = f a
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
i:nat
H9:(S i < S (Rlength l))%nat
Hreci:(i < S (Rlength l))%nat -> constant_D_eq {| fe := g'; pre := H8 |} (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i))
constant_D_eq {| fe := g'; pre := H8 |} (co_interval (pos_Rl (cons r (cons r1 l)) (S i)) (pos_Rl (cons r (cons r1 l)) (S (S i)))) (f (pos_Rl (cons r (cons r1 l)) (S i)))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:r = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
H9:(0 < S (Rlength l))%nat
x:R
H10:r <= x
H11:x < r1
r3:~ r1 <= x

f a = f a
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
i:nat
H9:(S i < S (Rlength l))%nat
Hreci:(i < S (Rlength l))%nat -> constant_D_eq {| fe := g'; pre := H8 |} (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i))
constant_D_eq {| fe := g'; pre := H8 |} (co_interval (pos_Rl (cons r (cons r1 l)) (S i)) (pos_Rl (cons r (cons r1 l)) (S (S i)))) (f (pos_Rl (cons r (cons r1 l)) (S i)))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
i:nat
H9:(S i < S (Rlength l))%nat
Hreci:(i < S (Rlength l))%nat -> constant_D_eq {| fe := g'; pre := H8 |} (co_interval (pos_Rl (cons r (cons r1 l)) i) (pos_Rl (cons r (cons r1 l)) (S i))) (f (pos_Rl (cons r (cons r1 l)) i))

constant_D_eq {| fe := g'; pre := H8 |} (co_interval (pos_Rl (cons r (cons r1 l)) (S i)) (pos_Rl (cons r (cons r1 l)) (S (S i)))) (f (pos_Rl (cons r (cons r1 l)) (S i)))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
i:nat
H9:(S i < S (Rlength l))%nat
H10:(i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))

(i < Init.Nat.pred (Rlength (cons r1 l)))%nat
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
i:nat
H9:(S i < S (Rlength l))%nat
H10:(i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
H11:(i < Init.Nat.pred (Rlength (cons r1 l)))%nat
constant_D_eq {| fe := g'; pre := H8 |} (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x : R => if Rle_dec r1 x then g x else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
i:nat
H9:(S i < S (Rlength l))%nat
H10:(i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
H11:(i < Init.Nat.pred (Rlength (cons r1 l)))%nat

constant_D_eq {| fe := g'; pre := H8 |} (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
i:nat
H9:(S i < S (Rlength l))%nat
H10:(i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
H11:(i < Init.Nat.pred (Rlength (cons r1 l)))%nat
H12:forall x0 : R, pos_Rl (cons r1 l) i <= x0 < pos_Rl (cons r1 l) (S i) -> g x0 = f (pos_Rl (cons r1 l) i)
x:R
H13:pos_Rl (cons r1 l) i <= x < pos_Rl (cons r1 l) (S i)
r0:r1 <= x

g x = g x
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
i:nat
H9:(S i < S (Rlength l))%nat
H10:(i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
H11:(i < Init.Nat.pred (Rlength (cons r1 l)))%nat
H12:forall x0 : R, pos_Rl (cons r1 l) i <= x0 < pos_Rl (cons r1 l) (S i) -> g x0 = f (pos_Rl (cons r1 l) i)
x:R
H13:pos_Rl (cons r1 l) i <= x < pos_Rl (cons r1 l) (S i)
r1 <= x
r, r1:R
l:Rlist
b:R
f:R -> R
IHl:forall a0 : R, ordered_Rlist (cons r1 l) -> pos_Rl (cons r1 l) 0 = a0 -> pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b -> {g0 : StepFun a0 b | g0 b = f b /\ (forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g0 (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0)))}
a:R
H:ordered_Rlist (cons r (cons r1 l))
H0:pos_Rl (cons r (cons r1 l)) 0 = a
H1:pos_Rl (cons r (cons r1 l)) (Init.Nat.pred (Rlength (cons r (cons r1 l)))) = b
H2:ordered_Rlist (cons r1 l)
H3:pos_Rl (cons r1 l) 0 = r1
H4:pos_Rl (cons r1 l) (Init.Nat.pred (Rlength (cons r1 l))) = b
g:StepFun r1 b
H5:g b = f b
H6:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i0) (pos_Rl (cons r1 l) (S i0))) (f (pos_Rl (cons r1 l) i0))
g':=fun x0 : R => if Rle_dec r1 x0 then g x0 else f a:R -> R
H7:r1 <= b
H8:IsStepFun g' a b
i:nat
H9:(S i < S (Rlength l))%nat
H10:(i < Init.Nat.pred (Rlength (cons r1 l)))%nat -> constant_D_eq g (co_interval (pos_Rl (cons r1 l) i) (pos_Rl (cons r1 l) (S i))) (f (pos_Rl (cons r1 l) i))
H11:(i < Init.Nat.pred (Rlength (cons r1 l)))%nat
H12:forall x0 : R, pos_Rl (cons r1 l) i <= x0 < pos_Rl (cons r1 l) (S i) -> g x0 = f (pos_Rl (cons r1 l) i)
x:R
H13:pos_Rl (cons r1 l) i <= x < pos_Rl (cons r1 l) (S i)

r1 <= x
elim H13; clear H13; intros; apply Rle_trans with (pos_Rl (cons r1 l) i); try assumption; change (pos_Rl (cons r1 l) 0 <= pos_Rl (cons r1 l) i); elim (RList_P6 (cons r1 l)); intros; apply H15; [ assumption | apply le_O_n | simpl; apply lt_trans with (Rlength l); [ apply lt_S_n; assumption | apply lt_n_Sn ] ]. Qed.

forall (a b : R) (f : StepFun a b), RiemannInt_SF f = - RiemannInt_SF {| fe := f; pre := StepFun_P6 (pre f) |}

forall (a b : R) (f : StepFun a b), RiemannInt_SF f = - RiemannInt_SF {| fe := f; pre := StepFun_P6 (pre f) |}
a, b:R
f:StepFun a b
r:b <= a
r0:a <= b

Int_SF (subdivision_val f) (subdivision f) = - Int_SF (subdivision_val {| fe := f; pre := StepFun_P6 (pre f) |}) (subdivision {| fe := f; pre := StepFun_P6 (pre f) |})
a, b:R
f:StepFun a b
n:~ b <= a
r:a <= b
Int_SF (subdivision_val f) (subdivision f) = - - Int_SF (subdivision_val {| fe := f; pre := StepFun_P6 (pre f) |}) (subdivision {| fe := f; pre := StepFun_P6 (pre f) |})
a, b:R
f:StepFun a b
r:b <= a
n:~ a <= b
- Int_SF (subdivision_val f) (subdivision f) = - Int_SF (subdivision_val {| fe := f; pre := StepFun_P6 (pre f) |}) (subdivision {| fe := f; pre := StepFun_P6 (pre f) |})
a, b:R
f:StepFun a b
n:~ b <= a
n0:~ a <= b
- Int_SF (subdivision_val f) (subdivision f) = - - Int_SF (subdivision_val {| fe := f; pre := StepFun_P6 (pre f) |}) (subdivision {| fe := f; pre := StepFun_P6 (pre f) |})
a, b:R
f:StepFun a b
n:~ b <= a
r:a <= b

Int_SF (subdivision_val f) (subdivision f) = - - Int_SF (subdivision_val {| fe := f; pre := StepFun_P6 (pre f) |}) (subdivision {| fe := f; pre := StepFun_P6 (pre f) |})
a, b:R
f:StepFun a b
r:b <= a
n:~ a <= b
- Int_SF (subdivision_val f) (subdivision f) = - Int_SF (subdivision_val {| fe := f; pre := StepFun_P6 (pre f) |}) (subdivision {| fe := f; pre := StepFun_P6 (pre f) |})
a, b:R
f:StepFun a b
n:~ b <= a
n0:~ a <= b
- Int_SF (subdivision_val f) (subdivision f) = - - Int_SF (subdivision_val {| fe := f; pre := StepFun_P6 (pre f) |}) (subdivision {| fe := f; pre := StepFun_P6 (pre f) |})
a, b:R
f:StepFun a b
r:b <= a
n:~ a <= b

- Int_SF (subdivision_val f) (subdivision f) = - Int_SF (subdivision_val {| fe := f; pre := StepFun_P6 (pre f) |}) (subdivision {| fe := f; pre := StepFun_P6 (pre f) |})
a, b:R
f:StepFun a b
n:~ b <= a
n0:~ a <= b
- Int_SF (subdivision_val f) (subdivision f) = - - Int_SF (subdivision_val {| fe := f; pre := StepFun_P6 (pre f) |}) (subdivision {| fe := f; pre := StepFun_P6 (pre f) |})
a, b:R
f:StepFun a b
n:~ b <= a
n0:~ a <= b

- Int_SF (subdivision_val f) (subdivision f) = - - Int_SF (subdivision_val {| fe := f; pre := StepFun_P6 (pre f) |}) (subdivision {| fe := f; pre := StepFun_P6 (pre f) |})
assert (H : a < b); [ auto with real | assert (H0 : b < a); [ auto with real | elim (Rlt_irrefl _ (Rlt_trans _ _ _ H H0)) ] ]. Qed.

forall (f : R -> R) (a b c : R) (l1 l2 lf1 lf2 : Rlist), a < b -> b < c -> adapted_couple f a b l1 lf1 -> adapted_couple f b c l2 lf2 -> adapted_couple f a c (cons_Rlist l1 l2) (FF (cons_Rlist l1 l2) f)

forall (f : R -> R) (a b c : R) (l1 l2 lf1 lf2 : Rlist), a < b -> b < c -> adapted_couple f a b l1 lf1 -> adapted_couple f b c l2 lf2 -> adapted_couple f a c (cons_Rlist l1 l2) (FF (cons_Rlist l1 l2) f)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)

ordered_Rlist (cons_Rlist l1 l2)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
pos_Rl (cons_Rlist l1 l2) 0 = Rmin a c
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
pos_Rl (cons_Rlist l1 l2) (Init.Nat.pred (Rlength (cons_Rlist l1 l2))) = Rmax a c
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
Rlength (cons_Rlist l1 l2) = S (Rlength (FF (cons_Rlist l1 l2) f))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_Rlist l1 l2) i) (pos_Rl (cons_Rlist l1 l2) (S i))) (pos_Rl (FF (cons_Rlist l1 l2) f) i)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)

pos_Rl l1 (Init.Nat.pred (Rlength l1)) <= pos_Rl l2 0
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
pos_Rl (cons_Rlist l1 l2) 0 = Rmin a c
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
pos_Rl (cons_Rlist l1 l2) (Init.Nat.pred (Rlength (cons_Rlist l1 l2))) = Rmax a c
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
Rlength (cons_Rlist l1 l2) = S (Rlength (FF (cons_Rlist l1 l2) f))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_Rlist l1 l2) i) (pos_Rl (cons_Rlist l1 l2) (S i))) (pos_Rl (FF (cons_Rlist l1 l2) f) i)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)

pos_Rl (cons_Rlist l1 l2) 0 = Rmin a c
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
pos_Rl (cons_Rlist l1 l2) (Init.Nat.pred (Rlength (cons_Rlist l1 l2))) = Rmax a c
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
Rlength (cons_Rlist l1 l2) = S (Rlength (FF (cons_Rlist l1 l2) f))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_Rlist l1 l2) i) (pos_Rl (cons_Rlist l1 l2) (S i))) (pos_Rl (FF (cons_Rlist l1 l2) f) i)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)

pos_Rl l1 0 = Rmin a c
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
l1 <> nil
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
pos_Rl (cons_Rlist l1 l2) (Init.Nat.pred (Rlength (cons_Rlist l1 l2))) = Rmax a c
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
Rlength (cons_Rlist l1 l2) = S (Rlength (FF (cons_Rlist l1 l2) f))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_Rlist l1 l2) i) (pos_Rl (cons_Rlist l1 l2) (S i))) (pos_Rl (FF (cons_Rlist l1 l2) f) i)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)

l1 <> nil
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
pos_Rl (cons_Rlist l1 l2) (Init.Nat.pred (Rlength (cons_Rlist l1 l2))) = Rmax a c
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
Rlength (cons_Rlist l1 l2) = S (Rlength (FF (cons_Rlist l1 l2) f))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_Rlist l1 l2) i) (pos_Rl (cons_Rlist l1 l2) (S i))) (pos_Rl (FF (cons_Rlist l1 l2) f) i)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)

pos_Rl (cons_Rlist l1 l2) (Init.Nat.pred (Rlength (cons_Rlist l1 l2))) = Rmax a c
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
Rlength (cons_Rlist l1 l2) = S (Rlength (FF (cons_Rlist l1 l2) f))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_Rlist l1 l2) i) (pos_Rl (cons_Rlist l1 l2) (S i))) (pos_Rl (FF (cons_Rlist l1 l2) f) i)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)

pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax a c
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
l2 <> nil
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
Rlength (cons_Rlist l1 l2) = S (Rlength (FF (cons_Rlist l1 l2) f))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_Rlist l1 l2) i) (pos_Rl (cons_Rlist l1 l2) (S i))) (pos_Rl (FF (cons_Rlist l1 l2) f) i)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)

l2 <> nil
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
Rlength (cons_Rlist l1 l2) = S (Rlength (FF (cons_Rlist l1 l2) f))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_Rlist l1 l2) i) (pos_Rl (cons_Rlist l1 l2) (S i))) (pos_Rl (FF (cons_Rlist l1 l2) f) i)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)

Rlength (cons_Rlist l1 l2) = S (Rlength (FF (cons_Rlist l1 l2) f))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_Rlist l1 l2) i) (pos_Rl (cons_Rlist l1 l2) (S i))) (pos_Rl (FF (cons_Rlist l1 l2) f) i)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)

(0 < Rlength (cons_Rlist l1 l2))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_Rlist l1 l2) i) (pos_Rl (cons_Rlist l1 l2) (S i))) (pos_Rl (FF (cons_Rlist l1 l2) f) i)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:0%nat = (Rlength l1 + Rlength l2)%nat

False
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_Rlist l1 l2) i) (pos_Rl (cons_Rlist l1 l2) (S i))) (pos_Rl (FF (cons_Rlist l1 l2) f) i)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:0%nat = (Rlength l1 + Rlength l2)%nat

(Rlength l1 + Rlength l2)%nat = 0%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:0%nat = (Rlength l1 + Rlength l2)%nat
H2:(Rlength l1 + Rlength l2)%nat = 0%nat
False
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_Rlist l1 l2) i) (pos_Rl (cons_Rlist l1 l2) (S i))) (pos_Rl (FF (cons_Rlist l1 l2) f) i)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:0%nat = (Rlength l1 + Rlength l2)%nat
H2:(Rlength l1 + Rlength l2)%nat = 0%nat

False
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_Rlist l1 l2) i) (pos_Rl (cons_Rlist l1 l2) (S i))) (pos_Rl (FF (cons_Rlist l1 l2) f) i)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)

forall i : nat, (i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons_Rlist l1 l2) i) (pos_Rl (cons_Rlist l1 l2) (S i))) (pos_Rl (FF (cons_Rlist l1 l2) f) i)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(S (S i) <= Rlength l1)%nat

f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(S (S i) <= Rlength l1)%nat

pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i

f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i

pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)

f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)

(2 <= Rlength l1)%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat

f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)

f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) i)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)

f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) 0)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
Hreci:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> pos_Rl l1 i < x < pos_Rl l1 (S i) -> (S (S i) <= Rlength l1)%nat -> pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i -> pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i) -> f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) i)
f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0

(0 < Init.Nat.pred (Rlength l1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
f x = f ((r1 + r2) / 2)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
Hreci:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> pos_Rl l1 i < x < pos_Rl l1 (S i) -> (S (S i) <= Rlength l1)%nat -> pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i -> pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i) -> f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) i)
f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat

f x = f ((r1 + r2) / 2)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
Hreci:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> pos_Rl l1 i < x < pos_Rl l1 (S i) -> (S (S i) <= Rlength l1)%nat -> pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i -> pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i) -> f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) i)
f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0

pos_Rl lf1 0 = pos_Rl lf1 0
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
pos_Rl l1 0 < (r1 + r2) / 2 < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
pos_Rl l1 0 < x < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
Hreci:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> pos_Rl l1 i < x < pos_Rl l1 (S i) -> (S (S i) <= Rlength l1)%nat -> pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i -> pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i) -> f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) i)
f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0

pos_Rl l1 0 < (r1 + r2) / 2 < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
pos_Rl l1 0 < x < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
Hreci:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> pos_Rl l1 i < x < pos_Rl l1 (S i) -> (S (S i) <= Rlength l1)%nat -> pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i -> pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i) -> f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) i)
f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0

r1 <= r2
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H21:r1 <= r2
pos_Rl l1 0 < (r1 + r2) / 2 < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
pos_Rl l1 0 < x < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
Hreci:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> pos_Rl l1 i < x < pos_Rl l1 (S i) -> (S (S i) <= Rlength l1)%nat -> pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i -> pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i) -> f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) i)
f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
r1, r2:R
r3:Rlist
H3:ordered_Rlist (cons r1 (cons r2 r3))
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0

(0 < Init.Nat.pred (Rlength (cons r1 (cons r2 r3))))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H21:r1 <= r2
pos_Rl l1 0 < (r1 + r2) / 2 < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
pos_Rl l1 0 < x < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
Hreci:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> pos_Rl l1 i < x < pos_Rl l1 (S i) -> (S (S i) <= Rlength l1)%nat -> pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i -> pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i) -> f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) i)
f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H21:r1 <= r2

pos_Rl l1 0 < (r1 + r2) / 2 < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
pos_Rl l1 0 < x < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
Hreci:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> pos_Rl l1 i < x < pos_Rl l1 (S i) -> (S (S i) <= Rlength l1)%nat -> pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i -> pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i) -> f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) i)
f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H21:r1 <= r2
H22:r1 < r2

pos_Rl l1 0 < (r1 + r2) / 2 < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H21:r1 <= r2
H22:r1 = r2
pos_Rl l1 0 < (r1 + r2) / 2 < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
pos_Rl l1 0 < x < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
Hreci:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> pos_Rl l1 i < x < pos_Rl l1 (S i) -> (S (S i) <= Rlength l1)%nat -> pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i -> pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i) -> f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) i)
f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H21:r1 <= r2
H22:r1 < r2

pos_Rl l1 0 < (r1 + r2) / 2
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H21:r1 <= r2
H22:r1 < r2
(r1 + r2) / 2 < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H21:r1 <= r2
H22:r1 = r2
pos_Rl l1 0 < (r1 + r2) / 2 < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
pos_Rl l1 0 < x < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
Hreci:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> pos_Rl l1 i < x < pos_Rl l1 (S i) -> (S (S i) <= Rlength l1)%nat -> pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i -> pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i) -> f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) i)
f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H21:r1 <= r2
H22:r1 < r2

(r1 + r2) / 2 < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H21:r1 <= r2
H22:r1 = r2
pos_Rl l1 0 < (r1 + r2) / 2 < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
pos_Rl l1 0 < x < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
Hreci:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> pos_Rl l1 i < x < pos_Rl l1 (S i) -> (S (S i) <= Rlength l1)%nat -> pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i -> pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i) -> f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) i)
f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H21:r1 <= r2
H22:r1 = r2

pos_Rl l1 0 < (r1 + r2) / 2 < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
pos_Rl l1 0 < x < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
Hreci:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> pos_Rl l1 i < x < pos_Rl l1 (S i) -> (S (S i) <= Rlength l1)%nat -> pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i -> pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i) -> f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) i)
f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 0 < x < pos_Rl l1 1
H12:(2 <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l1 0
H15:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l1 1
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
H18:(0 < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0
H19:(0 < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 0 < x0 < pos_Rl l1 1 -> f x0 = pos_Rl lf1 0

pos_Rl l1 0 < x < pos_Rl l1 1
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
Hreci:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> pos_Rl l1 i < x < pos_Rl l1 (S i) -> (S (S i) <= Rlength l1)%nat -> pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i -> pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i) -> f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) i)
f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
Hreci:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> pos_Rl l1 i < x < pos_Rl l1 (S i) -> (S (S i) <= Rlength l1)%nat -> pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i -> pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i) -> f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) i)

f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)

f x = f ((pos_Rl (cons_Rlist (cons r2 r3) l2) i + pos_Rl (cons_Rlist (cons r2 r3) l2) (S i)) / 2)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)

(S i < Init.Nat.pred (Rlength l1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
f x = f ((pos_Rl (cons r1 (cons r2 r3)) (S i) + pos_Rl (cons r1 (cons r2 r3)) (S (S i))) / 2)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat

f x = f ((pos_Rl (cons r1 (cons r2 r3)) (S i) + pos_Rl (cons r1 (cons r2 r3)) (S (S i))) / 2)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)

pos_Rl lf1 (S i) = pos_Rl lf1 (S i)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
pos_Rl l1 (S i) < (pos_Rl (cons r1 (cons r2 r3)) (S i) + pos_Rl (cons r1 (cons r2 r3)) (S (S i))) / 2 < pos_Rl l1 (S (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)

pos_Rl l1 (S i) < (pos_Rl (cons r1 (cons r2 r3)) (S i) + pos_Rl (cons r1 (cons r2 r3)) (S (S i))) / 2 < pos_Rl l1 (S (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)

pos_Rl l1 (S i) <= pos_Rl l1 (S (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H21:pos_Rl l1 (S i) <= pos_Rl l1 (S (S i))
pos_Rl l1 (S i) < (pos_Rl l1 (S i) + pos_Rl l1 (S (S i))) / 2 < pos_Rl l1 (S (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H21:pos_Rl l1 (S i) <= pos_Rl l1 (S (S i))

pos_Rl l1 (S i) < (pos_Rl l1 (S i) + pos_Rl l1 (S (S i))) / 2 < pos_Rl l1 (S (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H21:pos_Rl l1 (S i) <= pos_Rl l1 (S (S i))
H22:pos_Rl l1 (S i) < pos_Rl l1 (S (S i))

pos_Rl l1 (S i) < (pos_Rl l1 (S i) + pos_Rl l1 (S (S i))) / 2 < pos_Rl l1 (S (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H21:pos_Rl l1 (S i) <= pos_Rl l1 (S (S i))
H22:pos_Rl l1 (S i) = pos_Rl l1 (S (S i))
pos_Rl l1 (S i) < (pos_Rl l1 (S i) + pos_Rl l1 (S (S i))) / 2 < pos_Rl l1 (S (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H21:pos_Rl l1 (S i) <= pos_Rl l1 (S (S i))
H22:pos_Rl l1 (S i) < pos_Rl l1 (S (S i))

pos_Rl l1 (S i) < (pos_Rl l1 (S i) + pos_Rl l1 (S (S i))) / 2
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H21:pos_Rl l1 (S i) <= pos_Rl l1 (S (S i))
H22:pos_Rl l1 (S i) < pos_Rl l1 (S (S i))
(pos_Rl l1 (S i) + pos_Rl l1 (S (S i))) / 2 < pos_Rl l1 (S (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H21:pos_Rl l1 (S i) <= pos_Rl l1 (S (S i))
H22:pos_Rl l1 (S i) = pos_Rl l1 (S (S i))
pos_Rl l1 (S i) < (pos_Rl l1 (S i) + pos_Rl l1 (S (S i))) / 2 < pos_Rl l1 (S (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H21:pos_Rl l1 (S i) <= pos_Rl l1 (S (S i))
H22:pos_Rl l1 (S i) < pos_Rl l1 (S (S i))

(pos_Rl l1 (S i) + pos_Rl l1 (S (S i))) / 2 < pos_Rl l1 (S (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H21:pos_Rl l1 (S i) <= pos_Rl l1 (S (S i))
H22:pos_Rl l1 (S i) = pos_Rl l1 (S (S i))
pos_Rl l1 (S i) < (pos_Rl l1 (S i) + pos_Rl l1 (S (S i))) / 2 < pos_Rl l1 (S (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H21:pos_Rl l1 (S i) <= pos_Rl l1 (S (S i))
H22:pos_Rl l1 (S i) = pos_Rl l1 (S (S i))

pos_Rl l1 (S i) < (pos_Rl l1 (S i) + pos_Rl l1 (S (S i))) / 2 < pos_Rl l1 (S (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H14:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl (cons r1 (cons r2 r3)) (S i)
H15:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl (cons r1 (cons r2 r3)) (S (S i))
H16:(2 <= Rlength l1)%nat
H17:l1 = cons r1 (cons r2 r3)
H18:(S i < Init.Nat.pred (Rlength l1))%nat -> forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)
H19:(S i < Init.Nat.pred (Rlength l1))%nat
H20:forall x0 : R, pos_Rl l1 (S i) < x0 < pos_Rl l1 (S (S i)) -> f x0 = pos_Rl lf1 (S i)

pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 (S i) < x < pos_Rl l1 (S (S i))
H12:(S (S (S i)) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H15:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l1 (S (S i))
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)

(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl l1 i < x < pos_Rl l1 (S i)
H12:(S (S i) <= Rlength l1)%nat
H14:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l1 i
H15:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l1 (S i)
H16:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H17:l1 = cons r1 (cons r2 r3)

(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat

f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i

f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i

pos_Rl (cons_Rlist l1 l2) (S i) = b
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i
H16:pos_Rl (cons_Rlist l1 l2) (S i) = b
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i

pos_Rl l2 (S i - Rlength l1) = b
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i
(Rlength l1 <= S i)%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i
(S i < Rlength (cons_Rlist l1 l2))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i
H16:pos_Rl (cons_Rlist l1 l2) (S i) = b
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i

(Rlength l1 <= S i)%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i
(S i < Rlength (cons_Rlist l1 l2))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i
H16:pos_Rl (cons_Rlist l1 l2) (S i) = b
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i

(S i < Rlength (cons_Rlist l1 l2))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i
H16:pos_Rl (cons_Rlist l1 l2) (S i) = b
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist nil
H5:pos_Rl nil 0 = Rmin a b
H4:pos_Rl nil (Init.Nat.pred (Rlength nil)) = Rmax a b
H6:Rlength nil = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength nil))%nat -> constant_D_eq f (open_interval (pos_Rl nil i0) (pos_Rl nil (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist nil l2)))%nat
x:R
H2:pos_Rl (cons_Rlist nil l2) i < x < pos_Rl (cons_Rlist nil l2) (S i)
H12:(Rlength nil < S (S i))%nat
H15:Rlength nil = S i

(S i < Rlength (cons_Rlist nil l2))%nat
f:R -> R
a, b, c, r:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist (cons r l1)
H5:pos_Rl (cons r l1) 0 = Rmin a b
H4:pos_Rl (cons r l1) (Init.Nat.pred (Rlength (cons r l1))) = Rmax a b
H6:Rlength (cons r l1) = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r l1)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r l1) i0) (pos_Rl (cons r l1) (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist (cons r l1) l2)))%nat
x:R
H2:pos_Rl (cons_Rlist (cons r l1) l2) i < x < pos_Rl (cons_Rlist (cons r l1) l2) (S i)
H12:(Rlength (cons r l1) < S (S i))%nat
H15:Rlength (cons r l1) = S i
Hrecl1:ordered_Rlist l1 -> pos_Rl l1 0 = Rmin a b -> pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b -> Rlength l1 = S (Rlength lf1) -> (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)) -> (i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i) -> (Rlength l1 < S (S i))%nat -> Rlength l1 = S i -> (S i < Rlength (cons_Rlist l1 l2))%nat
(S i < Rlength (cons_Rlist (cons r l1) l2))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i
H16:pos_Rl (cons_Rlist l1 l2) (S i) = b
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c, r:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist (cons r l1)
H5:pos_Rl (cons r l1) 0 = Rmin a b
H4:pos_Rl (cons r l1) (Init.Nat.pred (Rlength (cons r l1))) = Rmax a b
H6:Rlength (cons r l1) = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r l1)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r l1) i0) (pos_Rl (cons r l1) (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist (cons r l1) l2)))%nat
x:R
H2:pos_Rl (cons_Rlist (cons r l1) l2) i < x < pos_Rl (cons_Rlist (cons r l1) l2) (S i)
H12:(Rlength (cons r l1) < S (S i))%nat
H15:Rlength (cons r l1) = S i
Hrecl1:ordered_Rlist l1 -> pos_Rl l1 0 = Rmin a b -> pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b -> Rlength l1 = S (Rlength lf1) -> (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)) -> (i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i) -> (Rlength l1 < S (S i))%nat -> Rlength l1 = S i -> (S i < Rlength (cons_Rlist l1 l2))%nat

(S i < Rlength (cons_Rlist (cons r l1) l2))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i
H16:pos_Rl (cons_Rlist l1 l2) (S i) = b
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i
H16:pos_Rl (cons_Rlist l1 l2) (S i) = b

f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i
H16:pos_Rl (cons_Rlist l1 l2) (S i) = b

pos_Rl (cons_Rlist l1 l2) i = b
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i
H16:pos_Rl (cons_Rlist l1 l2) (S i) = b
H17:pos_Rl (cons_Rlist l1 l2) i = b
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i
H16:pos_Rl (cons_Rlist l1 l2) (S i) = b

pos_Rl l1 i = b
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i
H16:pos_Rl (cons_Rlist l1 l2) (S i) = b
(i < Rlength l1)%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i
H16:pos_Rl (cons_Rlist l1 l2) (S i) = b
H17:pos_Rl (cons_Rlist l1 l2) i = b
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i
H16:pos_Rl (cons_Rlist l1 l2) (S i) = b

(i < Rlength l1)%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i
H16:pos_Rl (cons_Rlist l1 l2) (S i) = b
H17:pos_Rl (cons_Rlist l1 l2) i = b
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
H15:Rlength l1 = S i
H16:pos_Rl (cons_Rlist l1 l2) (S i) = b
H17:pos_Rl (cons_Rlist l1 l2) i = b

f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i

f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i

pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i

(Rlength l1 <= i)%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
(i < Rlength (cons_Rlist l1 l2))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i

(i < Rlength (cons_Rlist l1 l2))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)

f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)

pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)

pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
(S i - Rlength l1)%nat = S (i - Rlength l1)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)

(Rlength l1 <= S i)%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
(S i < Rlength (cons_Rlist l1 l2))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
(S i - Rlength l1)%nat = S (i - Rlength l1)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)

(S i < Rlength (cons_Rlist l1 l2))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
(S i - Rlength l1)%nat = S (i - Rlength l1)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist nil
H5:pos_Rl nil 0 = Rmin a b
H4:pos_Rl nil (Init.Nat.pred (Rlength nil)) = Rmax a b
H6:Rlength nil = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength nil))%nat -> constant_D_eq f (open_interval (pos_Rl nil i0) (pos_Rl nil (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist nil l2)))%nat
x:R
H2:pos_Rl (cons_Rlist nil l2) i < x < pos_Rl (cons_Rlist nil l2) (S i)
H12:(Rlength nil < S (S i))%nat
m:nat
H15:(S (Rlength nil) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist nil l2) i = pos_Rl l2 (i - Rlength nil)

(S i < Rlength (cons_Rlist nil l2))%nat
f:R -> R
a, b, c, r:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist (cons r l1)
H5:pos_Rl (cons r l1) 0 = Rmin a b
H4:pos_Rl (cons r l1) (Init.Nat.pred (Rlength (cons r l1))) = Rmax a b
H6:Rlength (cons r l1) = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r l1)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r l1) i0) (pos_Rl (cons r l1) (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist (cons r l1) l2)))%nat
x:R
H2:pos_Rl (cons_Rlist (cons r l1) l2) i < x < pos_Rl (cons_Rlist (cons r l1) l2) (S i)
H12:(Rlength (cons r l1) < S (S i))%nat
m:nat
H15:(S (Rlength (cons r l1)) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist (cons r l1) l2) i = pos_Rl l2 (i - Rlength (cons r l1))
Hrecl1:ordered_Rlist l1 -> pos_Rl l1 0 = Rmin a b -> pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b -> Rlength l1 = S (Rlength lf1) -> (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)) -> (i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i) -> (Rlength l1 < S (S i))%nat -> (S (Rlength l1) <= S i)%nat -> pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1) -> (S i < Rlength (cons_Rlist l1 l2))%nat
(S i < Rlength (cons_Rlist (cons r l1) l2))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
(S i - Rlength l1)%nat = S (i - Rlength l1)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c, r:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist (cons r l1)
H5:pos_Rl (cons r l1) 0 = Rmin a b
H4:pos_Rl (cons r l1) (Init.Nat.pred (Rlength (cons r l1))) = Rmax a b
H6:Rlength (cons r l1) = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r l1)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r l1) i0) (pos_Rl (cons r l1) (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist (cons r l1) l2)))%nat
x:R
H2:pos_Rl (cons_Rlist (cons r l1) l2) i < x < pos_Rl (cons_Rlist (cons r l1) l2) (S i)
H12:(Rlength (cons r l1) < S (S i))%nat
m:nat
H15:(S (Rlength (cons r l1)) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist (cons r l1) l2) i = pos_Rl l2 (i - Rlength (cons r l1))
Hrecl1:ordered_Rlist l1 -> pos_Rl l1 0 = Rmin a b -> pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b -> Rlength l1 = S (Rlength lf1) -> (forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)) -> (i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i) -> (Rlength l1 < S (S i))%nat -> (S (Rlength l1) <= S i)%nat -> pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1) -> (S i < Rlength (cons_Rlist l1 l2))%nat

(S i < Rlength (cons_Rlist (cons r l1) l2))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
(S i - Rlength l1)%nat = S (i - Rlength l1)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)

(S i - Rlength l1)%nat = S (i - Rlength l1)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))

f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))

(2 <= Rlength l1)%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
a, b:R
lf1:Rlist
H:a < b
H5:pos_Rl nil 0 = Rmin a b
H4:pos_Rl nil (Init.Nat.pred (Rlength nil)) = Rmax a b
H6:Rlength nil = S (Rlength lf1)

(2 <= Rlength nil)%nat
a, b, r:R
l1, lf1:Rlist
H:a < b
H5:pos_Rl (cons r l1) 0 = Rmin a b
H4:pos_Rl (cons r l1) (Init.Nat.pred (Rlength (cons r l1))) = Rmax a b
H6:Rlength (cons r l1) = S (Rlength lf1)
Hrecl1:pos_Rl l1 0 = Rmin a b -> pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b -> Rlength l1 = S (Rlength lf1) -> (2 <= Rlength l1)%nat
(2 <= Rlength (cons r l1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
a, b, r:R
l1, lf1:Rlist
H:a < b
H5:pos_Rl (cons r l1) 0 = Rmin a b
H4:pos_Rl (cons r l1) (Init.Nat.pred (Rlength (cons r l1))) = Rmax a b
H6:Rlength (cons r l1) = S (Rlength lf1)
Hrecl1:pos_Rl l1 0 = Rmin a b -> pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b -> Rlength l1 = S (Rlength lf1) -> (2 <= Rlength l1)%nat

(2 <= Rlength (cons r l1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
a, b, r:R
lf1:Rlist
H:a < b
H5:pos_Rl (cons r nil) 0 = Rmin a b
H4:pos_Rl (cons r nil) (Init.Nat.pred (Rlength (cons r nil))) = Rmax a b
H6:Rlength (cons r nil) = S (Rlength lf1)

(2 <= Rlength (cons r nil))%nat
a, b, r, r0:R
l1, lf1:Rlist
H:a < b
H5:pos_Rl (cons r (cons r0 l1)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r0 l1)) (Init.Nat.pred (Rlength (cons r (cons r0 l1)))) = Rmax a b
H6:Rlength (cons r (cons r0 l1)) = S (Rlength lf1)
Hrecl1:pos_Rl (cons r l1) 0 = Rmin a b -> pos_Rl (cons r l1) (Init.Nat.pred (Rlength (cons r l1))) = Rmax a b -> Rlength (cons r l1) = S (Rlength lf1) -> (2 <= Rlength (cons r l1))%nat
(2 <= Rlength (cons r (cons r0 l1)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
a, b, r:R
lf1:Rlist
H:a < b
H5:r = Rmin a b
H4:r = Rmax a b
H6:Rlength (cons r nil) = S (Rlength lf1)

Rmin a b < Rmax a b
a, b, r:R
lf1:Rlist
H:a < b
H5:r = Rmin a b
H4:r = Rmax a b
H6:Rlength (cons r nil) = S (Rlength lf1)
H0:Rmin a b < Rmax a b
(2 <= Rlength (cons r nil))%nat
a, b, r, r0:R
l1, lf1:Rlist
H:a < b
H5:pos_Rl (cons r (cons r0 l1)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r0 l1)) (Init.Nat.pred (Rlength (cons r (cons r0 l1)))) = Rmax a b
H6:Rlength (cons r (cons r0 l1)) = S (Rlength lf1)
Hrecl1:pos_Rl (cons r l1) 0 = Rmin a b -> pos_Rl (cons r l1) (Init.Nat.pred (Rlength (cons r l1))) = Rmax a b -> Rlength (cons r l1) = S (Rlength lf1) -> (2 <= Rlength (cons r l1))%nat
(2 <= Rlength (cons r (cons r0 l1)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
a, b, r:R
lf1:Rlist
H:a < b
H5:r = Rmin a b
H4:r = Rmax a b
H6:Rlength (cons r nil) = S (Rlength lf1)
H0:Rmin a b < Rmax a b

(2 <= Rlength (cons r nil))%nat
a, b, r, r0:R
l1, lf1:Rlist
H:a < b
H5:pos_Rl (cons r (cons r0 l1)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r0 l1)) (Init.Nat.pred (Rlength (cons r (cons r0 l1)))) = Rmax a b
H6:Rlength (cons r (cons r0 l1)) = S (Rlength lf1)
Hrecl1:pos_Rl (cons r l1) 0 = Rmin a b -> pos_Rl (cons r l1) (Init.Nat.pred (Rlength (cons r l1))) = Rmax a b -> Rlength (cons r l1) = S (Rlength lf1) -> (2 <= Rlength (cons r l1))%nat
(2 <= Rlength (cons r (cons r0 l1)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
a, b, r, r0:R
l1, lf1:Rlist
H:a < b
H5:pos_Rl (cons r (cons r0 l1)) 0 = Rmin a b
H4:pos_Rl (cons r (cons r0 l1)) (Init.Nat.pred (Rlength (cons r (cons r0 l1)))) = Rmax a b
H6:Rlength (cons r (cons r0 l1)) = S (Rlength lf1)
Hrecl1:pos_Rl (cons r l1) 0 = Rmin a b -> pos_Rl (cons r l1) (Init.Nat.pred (Rlength (cons r l1))) = Rmax a b -> Rlength (cons r l1) = S (Rlength lf1) -> (2 <= Rlength (cons r l1))%nat

(2 <= Rlength (cons r (cons r0 l1)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat

f x = pos_Rl (FF (cons_Rlist l1 l2) f) i
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)

f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) i)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i : nat, (i < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i) (pos_Rl l1 (S i))) (pos_Rl lf1 i)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H1:(0 < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) 0 < x < pos_Rl (cons_Rlist l1 l2) 1
H12:(Rlength l1 < 2)%nat
m:nat
H15:(S (Rlength l1) <= 1)%nat
H14:m = 1%nat
H16:pos_Rl (cons_Rlist l1 l2) 0 = pos_Rl l2 (0 - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) 1 = pos_Rl l2 (S (0 - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)

f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) 0)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
Hreci:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i) -> (Rlength l1 < S (S i))%nat -> (S (Rlength l1) <= S i)%nat -> m = S i -> pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1) -> pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1)) -> f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) i)
f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
Hreci:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat -> pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i) -> (Rlength l1 < S (S i))%nat -> (S (Rlength l1) <= S i)%nat -> m = S i -> pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1) -> pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1)) -> f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) i)

f x = f (pos_Rl (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1) (S i))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)

f x = f ((pos_Rl (cons_Rlist (cons r2 r3) l2) i + pos_Rl (cons_Rlist (cons r2 r3) l2) (S i)) / 2)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)

(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
f x = f ((pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3))) + pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))) / 2)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)

(S (S i) - Rlength l1 < Rlength l2)%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
(Rlength l1 <= S i)%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
f x = f ((pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3))) + pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))) / 2)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)

(S (S i) < Rlength l1 + Rlength l2)%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
(Rlength l1 <= S (S i))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
(Rlength l1 <= S i)%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
f x = f ((pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3))) + pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))) / 2)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)

(Rlength l1 <= S (S i))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
(Rlength l1 <= S i)%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
f x = f ((pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3))) + pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))) / 2)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)

(Rlength l1 <= S i)%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
f x = f ((pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3))) + pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))) / 2)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat

f x = f ((pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3))) + pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))) / 2)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)

pos_Rl lf2 (S i - Rlength l1) = pos_Rl lf2 (S i - Rlength l1)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
pos_Rl l2 (S i - Rlength l1) < (pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3))) + pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))) / 2 < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
pos_Rl l2 (S i - Rlength l1) < x < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)

pos_Rl l2 (S i - Rlength l1) < (pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3))) + pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))) / 2 < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
pos_Rl l2 (S i - Rlength l1) < x < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)

pos_Rl l2 (S i - Rlength l1) <= pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H23:pos_Rl l2 (S i - Rlength l1) <= pos_Rl l2 (S (S i - Rlength l1))
pos_Rl l2 (S i - Rlength l1) < (pos_Rl l2 (S i - Rlength l1) + pos_Rl l2 (S (S i - Rlength l1))) / 2 < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
pos_Rl l2 (S i - Rlength l1) < x < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)

(S (S i - Rlength l1) < Rlength l2)%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H23:pos_Rl l2 (S i - Rlength l1) <= pos_Rl l2 (S (S i - Rlength l1))
pos_Rl l2 (S i - Rlength l1) < (pos_Rl l2 (S i - Rlength l1) + pos_Rl l2 (S (S i - Rlength l1))) / 2 < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
pos_Rl l2 (S i - Rlength l1) < x < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)

(S (S i) - Rlength l1 < Rlength l2)%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
(Rlength l1 <= S i)%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H23:pos_Rl l2 (S i - Rlength l1) <= pos_Rl l2 (S (S i - Rlength l1))
pos_Rl l2 (S i - Rlength l1) < (pos_Rl l2 (S i - Rlength l1) + pos_Rl l2 (S (S i - Rlength l1))) / 2 < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
pos_Rl l2 (S i - Rlength l1) < x < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)

(S (S i) < Rlength l1 + Rlength l2)%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
(Rlength l1 <= S (S i))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
(Rlength l1 <= S i)%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H23:pos_Rl l2 (S i - Rlength l1) <= pos_Rl l2 (S (S i - Rlength l1))
pos_Rl l2 (S i - Rlength l1) < (pos_Rl l2 (S i - Rlength l1) + pos_Rl l2 (S (S i - Rlength l1))) / 2 < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
pos_Rl l2 (S i - Rlength l1) < x < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)

(Rlength l1 <= S (S i))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
(Rlength l1 <= S i)%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H23:pos_Rl l2 (S i - Rlength l1) <= pos_Rl l2 (S (S i - Rlength l1))
pos_Rl l2 (S i - Rlength l1) < (pos_Rl l2 (S i - Rlength l1) + pos_Rl l2 (S (S i - Rlength l1))) / 2 < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
pos_Rl l2 (S i - Rlength l1) < x < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)

(Rlength l1 <= S i)%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H23:pos_Rl l2 (S i - Rlength l1) <= pos_Rl l2 (S (S i - Rlength l1))
pos_Rl l2 (S i - Rlength l1) < (pos_Rl l2 (S i - Rlength l1) + pos_Rl l2 (S (S i - Rlength l1))) / 2 < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
pos_Rl l2 (S i - Rlength l1) < x < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H23:pos_Rl l2 (S i - Rlength l1) <= pos_Rl l2 (S (S i - Rlength l1))

pos_Rl l2 (S i - Rlength l1) < (pos_Rl l2 (S i - Rlength l1) + pos_Rl l2 (S (S i - Rlength l1))) / 2 < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
pos_Rl l2 (S i - Rlength l1) < x < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H23:pos_Rl l2 (S i - Rlength l1) <= pos_Rl l2 (S (S i - Rlength l1))
H24:pos_Rl l2 (S i - Rlength l1) < pos_Rl l2 (S (S i - Rlength l1))

pos_Rl l2 (S i - Rlength l1) < (pos_Rl l2 (S i - Rlength l1) + pos_Rl l2 (S (S i - Rlength l1))) / 2 < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H23:pos_Rl l2 (S i - Rlength l1) <= pos_Rl l2 (S (S i - Rlength l1))
H24:pos_Rl l2 (S i - Rlength l1) = pos_Rl l2 (S (S i - Rlength l1))
pos_Rl l2 (S i - Rlength l1) < (pos_Rl l2 (S i - Rlength l1) + pos_Rl l2 (S (S i - Rlength l1))) / 2 < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
pos_Rl l2 (S i - Rlength l1) < x < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H23:pos_Rl l2 (S i - Rlength l1) <= pos_Rl l2 (S (S i - Rlength l1))
H24:pos_Rl l2 (S i - Rlength l1) < pos_Rl l2 (S (S i - Rlength l1))

pos_Rl l2 (S i - Rlength l1) < (pos_Rl l2 (S i - Rlength l1) + pos_Rl l2 (S (S i - Rlength l1))) / 2
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H23:pos_Rl l2 (S i - Rlength l1) <= pos_Rl l2 (S (S i - Rlength l1))
H24:pos_Rl l2 (S i - Rlength l1) < pos_Rl l2 (S (S i - Rlength l1))
(pos_Rl l2 (S i - Rlength l1) + pos_Rl l2 (S (S i - Rlength l1))) / 2 < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H23:pos_Rl l2 (S i - Rlength l1) <= pos_Rl l2 (S (S i - Rlength l1))
H24:pos_Rl l2 (S i - Rlength l1) = pos_Rl l2 (S (S i - Rlength l1))
pos_Rl l2 (S i - Rlength l1) < (pos_Rl l2 (S i - Rlength l1) + pos_Rl l2 (S (S i - Rlength l1))) / 2 < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
pos_Rl l2 (S i - Rlength l1) < x < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H23:pos_Rl l2 (S i - Rlength l1) <= pos_Rl l2 (S (S i - Rlength l1))
H24:pos_Rl l2 (S i - Rlength l1) < pos_Rl l2 (S (S i - Rlength l1))

(pos_Rl l2 (S i - Rlength l1) + pos_Rl l2 (S (S i - Rlength l1))) / 2 < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H23:pos_Rl l2 (S i - Rlength l1) <= pos_Rl l2 (S (S i - Rlength l1))
H24:pos_Rl l2 (S i - Rlength l1) = pos_Rl l2 (S (S i - Rlength l1))
pos_Rl l2 (S i - Rlength l1) < (pos_Rl l2 (S i - Rlength l1) + pos_Rl l2 (S (S i - Rlength l1))) / 2 < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
pos_Rl l2 (S i - Rlength l1) < x < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H23:pos_Rl l2 (S i - Rlength l1) <= pos_Rl l2 (S (S i - Rlength l1))
H24:pos_Rl l2 (S i - Rlength l1) = pos_Rl l2 (S (S i - Rlength l1))

pos_Rl l2 (S i - Rlength l1) < (pos_Rl l2 (S i - Rlength l1) + pos_Rl l2 (S (S i - Rlength l1))) / 2 < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
pos_Rl l2 (S i - Rlength l1) < x < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)

pos_Rl l2 (S i - Rlength l1) < x < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)

pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H23:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
pos_Rl l2 (S i - Rlength l1) < x < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H23:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)

pos_Rl l2 (S i - Rlength l1) < x < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H23:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)

pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H23:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H24:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
pos_Rl l2 (S i - Rlength l1) < x < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
r1, r2:R
r3:Rlist
H16:pos_Rl (cons_Rlist (cons r2 r3) l2) i = pos_Rl l2 (S i - Rlength (cons r1 (cons r2 r3)))
H17:pos_Rl (cons_Rlist (cons r2 r3) l2) (S i) = pos_Rl l2 (S (S i - Rlength (cons r1 (cons r2 r3))))
H18:(2 <= Rlength l1)%nat
H19:l1 = cons r1 (cons r2 r3)
H20:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat -> forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H21:(S i - Rlength l1 < Init.Nat.pred (Rlength l2))%nat
H22:forall x0 : R, pos_Rl l2 (S i - Rlength l1) < x0 < pos_Rl l2 (S (S i - Rlength l1)) -> f x0 = pos_Rl lf2 (S i - Rlength l1)
H23:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H24:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))

pos_Rl l2 (S i - Rlength l1) < x < pos_Rl l2 (S (S i - Rlength l1))
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(S i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) (S i) < x < pos_Rl (cons_Rlist l1 l2) (S (S i))
H12:(Rlength l1 < S (S (S i)))%nat
m:nat
H15:(S (Rlength l1) <= S (S i))%nat
H14:m = S (S i)
H16:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S (S i)) = pos_Rl l2 (S (S i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)

(i < Init.Nat.pred (Rlength (cons_Rlist (cons r2 r3) l2)))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)
(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
f:R -> R
a, b, c:R
l1, l2, lf1, lf2:Rlist
H:a < b
H0:b < c
H3:ordered_Rlist l1
H5:pos_Rl l1 0 = Rmin a b
H4:pos_Rl l1 (Init.Nat.pred (Rlength l1)) = Rmax a b
H6:Rlength l1 = S (Rlength lf1)
H8:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l1))%nat -> constant_D_eq f (open_interval (pos_Rl l1 i0) (pos_Rl l1 (S i0))) (pos_Rl lf1 i0)
H7:ordered_Rlist l2
H10:pos_Rl l2 0 = Rmin b c
H9:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H11:Rlength l2 = S (Rlength lf2)
H13:forall i0 : nat, (i0 < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i0) (pos_Rl l2 (S i0))) (pos_Rl lf2 i0)
i:nat
H1:(i < Init.Nat.pred (Rlength (cons_Rlist l1 l2)))%nat
x:R
H2:pos_Rl (cons_Rlist l1 l2) i < x < pos_Rl (cons_Rlist l1 l2) (S i)
H12:(Rlength l1 < S (S i))%nat
m:nat
H15:(S (Rlength l1) <= S i)%nat
H14:m = S i
H16:pos_Rl (cons_Rlist l1 l2) i = pos_Rl l2 (i - Rlength l1)
H17:pos_Rl (cons_Rlist l1 l2) (S i) = pos_Rl l2 (S (i - Rlength l1))
H18:(2 <= Rlength l1)%nat
r1, r2:R
r3:Rlist
H19:l1 = cons r1 (cons r2 r3)

(i < Rlength (mid_Rlist (cons_Rlist (cons r2 r3) l2) r1))%nat
rewrite RList_P14; rewrite H19 in H1; simpl in H1; simpl; apply H1. Qed.

forall (f : R -> R) (a b c : R), a <= b -> b <= c -> IsStepFun f a b -> IsStepFun f b c -> IsStepFun f a c

forall (f : R -> R) (a b c : R), a <= b -> b <= c -> IsStepFun f a b -> IsStepFun f b c -> IsStepFun f a c
f:R -> R
a, b, c:R
H:a <= b
H0:b <= c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
Hltab:a < b

IsStepFun f a c
f:R -> R
a, b, c:R
H:a <= b
H0:b <= c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
Hab:a = b
IsStepFun f a c
f:R -> R
a, b, c:R
H:a <= b
H0:b <= c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
Hgtab:a > b
IsStepFun f a c
f:R -> R
a, b, c:R
H:a <= b
H0:b <= c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
Hltab:a < b
Hltbc:b < c

IsStepFun f a c
f:R -> R
a, b, c:R
H:a <= b
H0:b <= c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
Hltab:a < b
Hbc:b = c
IsStepFun f a c
f:R -> R
a, b, c:R
H:a <= b
H0:b <= c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
Hltab:a < b
Hgtbc:b > c
IsStepFun f a c
f:R -> R
a, b, c:R
H:a <= b
H0:b <= c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
Hab:a = b
IsStepFun f a c
f:R -> R
a, b, c:R
H:a <= b
H0:b <= c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
Hgtab:a > b
IsStepFun f a c
f:R -> R
a, b, c:R
H:a <= b
H0:b <= c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
Hltab:a < b
Hbc:b = c

IsStepFun f a c
f:R -> R
a, b, c:R
H:a <= b
H0:b <= c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
Hltab:a < b
Hgtbc:b > c
IsStepFun f a c
f:R -> R
a, b, c:R
H:a <= b
H0:b <= c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
Hab:a = b
IsStepFun f a c
f:R -> R
a, b, c:R
H:a <= b
H0:b <= c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
Hgtab:a > b
IsStepFun f a c
f:R -> R
a, b, c:R
H:a <= b
H0:b <= c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
Hltab:a < b
Hgtbc:b > c

IsStepFun f a c
f:R -> R
a, b, c:R
H:a <= b
H0:b <= c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
Hab:a = b
IsStepFun f a c
f:R -> R
a, b, c:R
H:a <= b
H0:b <= c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
Hgtab:a > b
IsStepFun f a c
f:R -> R
a, b, c:R
H:a <= b
H0:b <= c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
Hab:a = b

IsStepFun f a c
f:R -> R
a, b, c:R
H:a <= b
H0:b <= c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
Hgtab:a > b
IsStepFun f a c
f:R -> R
a, b, c:R
H:a <= b
H0:b <= c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
Hgtab:a > b

IsStepFun f a c
elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H Hgtab)). Qed.

forall (l1 l2 : Rlist) (f : R -> R), pos_Rl l1 (Init.Nat.pred (Rlength l1)) = pos_Rl l2 0 -> Int_SF (FF (cons_Rlist l1 l2) f) (cons_Rlist l1 l2) = Int_SF (FF l1 f) l1 + Int_SF (FF l2 f) l2

forall (l1 l2 : Rlist) (f : R -> R), pos_Rl l1 (Init.Nat.pred (Rlength l1)) = pos_Rl l2 0 -> Int_SF (FF (cons_Rlist l1 l2) f) (cons_Rlist l1 l2) = Int_SF (FF l1 f) l1 + Int_SF (FF l2 f) l2
intros l1 l2 f; induction l1 as [| r l1 IHl1]; intros H; [ simpl; ring | destruct l1 as [| r0 r1]; [ simpl in H; simpl; destruct l2 as [| r0 r1]; [ simpl; ring | simpl; simpl in H; rewrite H; ring ] | simpl; rewrite Rplus_assoc; apply Rplus_eq_compat_l; apply IHl1; rewrite <- H; reflexivity ] ]. Qed.

forall (f : R -> R) (a b c : R) (pr1 : IsStepFun f a b) (pr2 : IsStepFun f b c) (pr3 : IsStepFun f a c), RiemannInt_SF {| fe := f; pre := pr1 |} + RiemannInt_SF {| fe := f; pre := pr2 |} = RiemannInt_SF {| fe := f; pre := pr3 |}

forall (f : R -> R) (a b c : R) (pr1 : IsStepFun f a b) (pr2 : IsStepFun f b c) (pr3 : IsStepFun f a c), RiemannInt_SF {| fe := f; pre := pr1 |} + RiemannInt_SF {| fe := f; pre := pr2 |} = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c

RiemannInt_SF {| fe := f; pre := pr1 |} + RiemannInt_SF {| fe := f; pre := pr2 |} = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1

RiemannInt_SF {| fe := f; pre := pr1 |} + RiemannInt_SF {| fe := f; pre := pr2 |} = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2

RiemannInt_SF {| fe := f; pre := pr1 |} + RiemannInt_SF {| fe := f; pre := pr2 |} = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3

RiemannInt_SF {| fe := f; pre := pr1 |} + RiemannInt_SF {| fe := f; pre := pr2 |} = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3

(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) + RiemannInt_SF {| fe := f; pre := pr2 |} = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3

(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) + (if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3

(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) + (if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = (if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c

Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b

Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a = b
Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b < c

Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b = c
Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a = b
Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b < c

Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF (FF (cons_Rlist l1 l2) f) (cons_Rlist l1 l2)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b < c
Int_SF (FF (cons_Rlist l1 l2) f) (cons_Rlist l1 l2) = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b = c
Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a = b
Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b < c

Int_SF (FF l1 f) l1 + Int_SF lf2 l2 = Int_SF (FF (cons_Rlist l1 l2) f) (cons_Rlist l1 l2)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b < c
Int_SF (FF l1 f) l1 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b < c
Int_SF (FF (cons_Rlist l1 l2) f) (cons_Rlist l1 l2) = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b = c
Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a = b
Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b < c

Int_SF (FF l1 f) l1 + Int_SF (FF l2 f) l2 = Int_SF (FF (cons_Rlist l1 l2) f) (cons_Rlist l1 l2)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b < c
Int_SF (FF l2 f) l2 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b < c
Int_SF (FF l1 f) l1 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b < c
Int_SF (FF (cons_Rlist l1 l2) f) (cons_Rlist l1 l2) = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b = c
Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a = b
Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b < c

pos_Rl l1 (Init.Nat.pred (Rlength l1)) = pos_Rl l2 0
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b < c
Int_SF (FF l2 f) l2 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b < c
Int_SF (FF l1 f) l1 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b < c
Int_SF (FF (cons_Rlist l1 l2) f) (cons_Rlist l1 l2) = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b = c
Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a = b
Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b < c

Int_SF (FF l2 f) l2 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b < c
Int_SF (FF l1 f) l1 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b < c
Int_SF (FF (cons_Rlist l1 l2) f) (cons_Rlist l1 l2) = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b = c
Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a = b
Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b < c

Int_SF (FF l1 f) l1 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b < c
Int_SF (FF (cons_Rlist l1 l2) f) (cons_Rlist l1 l2) = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b = c
Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a = b
Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b < c

Int_SF (FF (cons_Rlist l1 l2) f) (cons_Rlist l1 l2) = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b = c
Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a = b
Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b = c

Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a = b
Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b = c

Int_SF lf1 l1 + 0 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b = c
0 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a = b
Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a < b
H0:b = c

0 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a = b
Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a = b

Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a = b

0 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a = b
0 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hle'':a <= c
H:a = b

0 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hle':b <= c
Hnle'':~ a <= c

Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c

Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c

Int_SF lf1 l1 = Int_SF lf2 l2 + Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c

c < b
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
Int_SF lf1 l1 = Int_SF lf2 l2 + Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b

Int_SF lf1 l1 = Int_SF lf2 l2 + Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c

Int_SF lf1 l1 = Int_SF lf2 l2 + Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a = c
Int_SF lf1 l1 = Int_SF lf2 l2 + Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c

Int_SF (FF (cons_Rlist l3 l2) f) (cons_Rlist l3 l2) = Int_SF lf3 l3 + Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c
Int_SF (FF (cons_Rlist l3 l2) f) (cons_Rlist l3 l2) = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a = c
Int_SF lf1 l1 = Int_SF lf2 l2 + Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c

Int_SF (FF (cons_Rlist l3 l2) f) (cons_Rlist l3 l2) = Int_SF (FF l3 f) l3 + Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c
Int_SF (FF l3 f) l3 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c
Int_SF (FF (cons_Rlist l3 l2) f) (cons_Rlist l3 l2) = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a = c
Int_SF lf1 l1 = Int_SF lf2 l2 + Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c

Int_SF (FF (cons_Rlist l3 l2) f) (cons_Rlist l3 l2) = Int_SF (FF l3 f) l3 + Int_SF (FF l2 f) l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c
Int_SF (FF l2 f) l2 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c
Int_SF (FF l3 f) l3 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c
Int_SF (FF (cons_Rlist l3 l2) f) (cons_Rlist l3 l2) = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a = c
Int_SF lf1 l1 = Int_SF lf2 l2 + Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c

pos_Rl l3 (Init.Nat.pred (Rlength l3)) = pos_Rl l2 0
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c
Int_SF (FF l2 f) l2 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c
Int_SF (FF l3 f) l3 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c
Int_SF (FF (cons_Rlist l3 l2) f) (cons_Rlist l3 l2) = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a = c
Int_SF lf1 l1 = Int_SF lf2 l2 + Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2, l3, lf3:Rlist
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c
H4:ordered_Rlist l2
H6:pos_Rl l2 0 = Rmin b c
H5:pos_Rl l2 (Init.Nat.pred (Rlength l2)) = Rmax b c
H7:Rlength l2 = S (Rlength lf2)
H9:forall i : nat, (i < Init.Nat.pred (Rlength l2))%nat -> constant_D_eq f (open_interval (pos_Rl l2 i) (pos_Rl l2 (S i))) (pos_Rl lf2 i)
H8:ordered_Rlist l3
H11:pos_Rl l3 0 = Rmin a c
H10:pos_Rl l3 (Init.Nat.pred (Rlength l3)) = Rmax a c
H12:Rlength l3 = S (Rlength lf3)
H14:forall i : nat, (i < Init.Nat.pred (Rlength l3))%nat -> constant_D_eq f (open_interval (pos_Rl l3 i) (pos_Rl l3 (S i))) (pos_Rl lf3 i)

(if Rle_dec a c then c else a) = (if Rle_dec b c then b else c)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c
Int_SF (FF l2 f) l2 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c
Int_SF (FF l3 f) l3 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c
Int_SF (FF (cons_Rlist l3 l2) f) (cons_Rlist l3 l2) = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a = c
Int_SF lf1 l1 = Int_SF lf2 l2 + Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c

Int_SF (FF l2 f) l2 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c
Int_SF (FF l3 f) l3 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c
Int_SF (FF (cons_Rlist l3 l2) f) (cons_Rlist l3 l2) = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a = c
Int_SF lf1 l1 = Int_SF lf2 l2 + Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c

Int_SF (FF l3 f) l3 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c
Int_SF (FF (cons_Rlist l3 l2) f) (cons_Rlist l3 l2) = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a = c
Int_SF lf1 l1 = Int_SF lf2 l2 + Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a < c

Int_SF (FF (cons_Rlist l3 l2) f) (cons_Rlist l3 l2) = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a = c
Int_SF lf1 l1 = Int_SF lf2 l2 + Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a = c

Int_SF lf1 l1 = Int_SF lf2 l2 + Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a = c

Int_SF lf1 l1 = Int_SF lf2 l2 + 0
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a = c
0 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hle'':a <= c
H:c < b
H0:a = c

0 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c

Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c

Int_SF lf1 l1 + - (Int_SF lf3 l3 + Int_SF lf1 l1) = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
Int_SF lf3 l3 + Int_SF lf1 l1 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c

Int_SF lf3 l3 + Int_SF lf1 l1 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b

Int_SF lf3 l3 + Int_SF lf1 l1 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a = b
Int_SF lf3 l3 + Int_SF lf1 l1 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b

Int_SF lf3 l3 + Int_SF lf1 l1 = Int_SF (FF (cons_Rlist l3 l1) f) (cons_Rlist l3 l1)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b
Int_SF (FF (cons_Rlist l3 l1) f) (cons_Rlist l3 l1) = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a = b
Int_SF lf3 l3 + Int_SF lf1 l1 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b

Int_SF (FF l3 f) l3 + Int_SF lf1 l1 = Int_SF (FF (cons_Rlist l3 l1) f) (cons_Rlist l3 l1)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b
Int_SF (FF l3 f) l3 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b
Int_SF (FF (cons_Rlist l3 l1) f) (cons_Rlist l3 l1) = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a = b
Int_SF lf3 l3 + Int_SF lf1 l1 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b

Int_SF (FF l3 f) l3 + Int_SF (FF l1 f) l1 = Int_SF (FF (cons_Rlist l3 l1) f) (cons_Rlist l3 l1)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b
Int_SF (FF l1 f) l1 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b
Int_SF (FF l3 f) l3 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b
Int_SF (FF (cons_Rlist l3 l1) f) (cons_Rlist l3 l1) = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a = b
Int_SF lf3 l3 + Int_SF lf1 l1 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b

pos_Rl l3 (Init.Nat.pred (Rlength l3)) = pos_Rl l1 0
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b
Int_SF (FF l1 f) l1 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b
Int_SF (FF l3 f) l3 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b
Int_SF (FF (cons_Rlist l3 l1) f) (cons_Rlist l3 l1) = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a = b
Int_SF lf3 l3 + Int_SF lf1 l1 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b

Int_SF (FF l1 f) l1 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b
Int_SF (FF l3 f) l3 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b
Int_SF (FF (cons_Rlist l3 l1) f) (cons_Rlist l3 l1) = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a = b
Int_SF lf3 l3 + Int_SF lf1 l1 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b

Int_SF (FF l3 f) l3 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b
Int_SF (FF (cons_Rlist l3 l1) f) (cons_Rlist l3 l1) = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a = b
Int_SF lf3 l3 + Int_SF lf1 l1 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b

Int_SF (FF (cons_Rlist l3 l1) f) (cons_Rlist l3 l1) = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a = b
Int_SF lf3 l3 + Int_SF lf1 l1 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b

adapted_couple ?f ?a ?b (cons_Rlist l3 l1) (FF (cons_Rlist l3 l1) f)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b
adapted_couple ?f ?a ?b l2 lf2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a = b
Int_SF lf3 l3 + Int_SF lf1 l1 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b

c < a
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b
H0:c < a
adapted_couple ?f ?a ?b (cons_Rlist l3 l1) (FF (cons_Rlist l3 l1) f)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b
adapted_couple ?f ?a ?b l2 lf2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a = b
Int_SF lf3 l3 + Int_SF lf1 l1 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b
H0:c < a

adapted_couple ?f ?a ?b (cons_Rlist l3 l1) (FF (cons_Rlist l3 l1) f)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b
adapted_couple ?f ?a ?b l2 lf2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a = b
Int_SF lf3 l3 + Int_SF lf1 l1 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a < b

adapted_couple f c b l2 lf2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a = b
Int_SF lf3 l3 + Int_SF lf1 l1 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a = b

Int_SF lf3 l3 + Int_SF lf1 l1 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a = b

Int_SF lf3 l3 + 0 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a = b
0 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hle:a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:a = b

0 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c

- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c

b < a
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a

- Int_SF lf1 l1 + Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a

- Int_SF lf1 l1 + (Int_SF lf3 l3 + Int_SF lf1 l1) = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
Int_SF lf3 l3 + Int_SF lf1 l1 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a

Int_SF lf3 l3 + Int_SF lf1 l1 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c

Int_SF lf1 l1 + Int_SF lf3 l3 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a = c
Int_SF lf1 l1 + Int_SF lf3 l3 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c

Int_SF lf1 l1 + Int_SF lf3 l3 = Int_SF (FF (cons_Rlist l1 l3) f) (cons_Rlist l1 l3)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c
Int_SF (FF (cons_Rlist l1 l3) f) (cons_Rlist l1 l3) = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a = c
Int_SF lf1 l1 + Int_SF lf3 l3 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c

Int_SF lf1 l1 + Int_SF (FF l3 f) l3 = Int_SF (FF (cons_Rlist l1 l3) f) (cons_Rlist l1 l3)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c
Int_SF (FF l3 f) l3 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c
Int_SF (FF (cons_Rlist l1 l3) f) (cons_Rlist l1 l3) = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a = c
Int_SF lf1 l1 + Int_SF lf3 l3 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c

Int_SF (FF l1 f) l1 + Int_SF (FF l3 f) l3 = Int_SF (FF (cons_Rlist l1 l3) f) (cons_Rlist l1 l3)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c
Int_SF (FF l1 f) l1 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c
Int_SF (FF l3 f) l3 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c
Int_SF (FF (cons_Rlist l1 l3) f) (cons_Rlist l1 l3) = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a = c
Int_SF lf1 l1 + Int_SF lf3 l3 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c

pos_Rl l1 (Init.Nat.pred (Rlength l1)) = pos_Rl l3 0
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c
Int_SF (FF l1 f) l1 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c
Int_SF (FF l3 f) l3 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c
Int_SF (FF (cons_Rlist l1 l3) f) (cons_Rlist l1 l3) = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a = c
Int_SF lf1 l1 + Int_SF lf3 l3 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c

Int_SF (FF l1 f) l1 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c
Int_SF (FF l3 f) l3 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c
Int_SF (FF (cons_Rlist l1 l3) f) (cons_Rlist l1 l3) = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a = c
Int_SF lf1 l1 + Int_SF lf3 l3 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c

Int_SF (FF l3 f) l3 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c
Int_SF (FF (cons_Rlist l1 l3) f) (cons_Rlist l1 l3) = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a = c
Int_SF lf1 l1 + Int_SF lf3 l3 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c

Int_SF (FF (cons_Rlist l1 l3) f) (cons_Rlist l1 l3) = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a = c
Int_SF lf1 l1 + Int_SF lf3 l3 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c

adapted_couple ?f ?a ?b (cons_Rlist l1 l3) (FF (cons_Rlist l1 l3) f)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c
adapted_couple ?f ?a ?b l2 lf2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a = c
Int_SF lf1 l1 + Int_SF lf3 l3 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a < c

adapted_couple f b c l2 lf2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a = c
Int_SF lf1 l1 + Int_SF lf3 l3 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a = c

Int_SF lf1 l1 + Int_SF lf3 l3 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a = c

Int_SF lf1 l1 + 0 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a = c
0 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hle'':a <= c
H:b < a
H0:a = c

0 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c

- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c

c < a
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a

- Int_SF lf1 l1 + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a

- (Int_SF lf2 l2 + Int_SF lf3 l3) + Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
Int_SF lf2 l2 + Int_SF lf3 l3 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a

Int_SF lf2 l2 + Int_SF lf3 l3 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c

Int_SF lf2 l2 + Int_SF lf3 l3 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b = c
Int_SF lf2 l2 + Int_SF lf3 l3 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c

Int_SF lf2 l2 + Int_SF lf3 l3 = Int_SF (FF (cons_Rlist l2 l3) f) (cons_Rlist l2 l3)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c
Int_SF (FF (cons_Rlist l2 l3) f) (cons_Rlist l2 l3) = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b = c
Int_SF lf2 l2 + Int_SF lf3 l3 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c

Int_SF lf2 l2 + Int_SF (FF l3 f) l3 = Int_SF (FF (cons_Rlist l2 l3) f) (cons_Rlist l2 l3)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c
Int_SF (FF l3 f) l3 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c
Int_SF (FF (cons_Rlist l2 l3) f) (cons_Rlist l2 l3) = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b = c
Int_SF lf2 l2 + Int_SF lf3 l3 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c

Int_SF (FF l2 f) l2 + Int_SF (FF l3 f) l3 = Int_SF (FF (cons_Rlist l2 l3) f) (cons_Rlist l2 l3)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c
Int_SF (FF l2 f) l2 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c
Int_SF (FF l3 f) l3 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c
Int_SF (FF (cons_Rlist l2 l3) f) (cons_Rlist l2 l3) = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b = c
Int_SF lf2 l2 + Int_SF lf3 l3 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c

pos_Rl l2 (Init.Nat.pred (Rlength l2)) = pos_Rl l3 0
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c
Int_SF (FF l2 f) l2 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c
Int_SF (FF l3 f) l3 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c
Int_SF (FF (cons_Rlist l2 l3) f) (cons_Rlist l2 l3) = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b = c
Int_SF lf2 l2 + Int_SF lf3 l3 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c

Int_SF (FF l2 f) l2 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c
Int_SF (FF l3 f) l3 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c
Int_SF (FF (cons_Rlist l2 l3) f) (cons_Rlist l2 l3) = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b = c
Int_SF lf2 l2 + Int_SF lf3 l3 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c

Int_SF (FF l3 f) l3 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c
Int_SF (FF (cons_Rlist l2 l3) f) (cons_Rlist l2 l3) = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b = c
Int_SF lf2 l2 + Int_SF lf3 l3 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c

Int_SF (FF (cons_Rlist l2 l3) f) (cons_Rlist l2 l3) = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b = c
Int_SF lf2 l2 + Int_SF lf3 l3 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c

adapted_couple ?f ?a ?b (cons_Rlist l2 l3) (FF (cons_Rlist l2 l3) f)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c
adapted_couple ?f ?a ?b l1 lf1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b = c
Int_SF lf2 l2 + Int_SF lf3 l3 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b < c

adapted_couple f b a l1 lf1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b = c
Int_SF lf2 l2 + Int_SF lf3 l3 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b = c

Int_SF lf2 l2 + Int_SF lf3 l3 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b = c

0 + Int_SF lf3 l3 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b = c
0 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hle':b <= c
Hnle'':~ a <= c
H:c < a
H0:b = c

0 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c

- Int_SF lf1 l1 + - Int_SF lf2 l2 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hle'':a <= c

b <= a
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c

- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c

c < b
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b

- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b

b < a
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a
- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a

- Int_SF lf1 l1 + - Int_SF lf2 l2 = - Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a

- Int_SF lf1 l1 + - Int_SF lf2 l2 = - (Int_SF lf2 l2 + Int_SF lf1 l1)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a
Int_SF lf2 l2 + Int_SF lf1 l1 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a

Int_SF lf2 l2 + Int_SF lf1 l1 = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a

Int_SF lf2 l2 + Int_SF lf1 l1 = Int_SF (FF (cons_Rlist l2 l1) f) (cons_Rlist l2 l1)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a
Int_SF (FF (cons_Rlist l2 l1) f) (cons_Rlist l2 l1) = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a

Int_SF lf2 l2 + Int_SF (FF l1 f) l1 = Int_SF (FF (cons_Rlist l2 l1) f) (cons_Rlist l2 l1)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a
Int_SF (FF l1 f) l1 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a
Int_SF (FF (cons_Rlist l2 l1) f) (cons_Rlist l2 l1) = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a

Int_SF (FF l2 f) l2 + Int_SF (FF l1 f) l1 = Int_SF (FF (cons_Rlist l2 l1) f) (cons_Rlist l2 l1)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a
Int_SF (FF l2 f) l2 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a
Int_SF (FF l1 f) l1 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a
Int_SF (FF (cons_Rlist l2 l1) f) (cons_Rlist l2 l1) = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a

pos_Rl l2 (Init.Nat.pred (Rlength l2)) = pos_Rl l1 0
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a
Int_SF (FF l2 f) l2 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a
Int_SF (FF l1 f) l1 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a
Int_SF (FF (cons_Rlist l2 l1) f) (cons_Rlist l2 l1) = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a

Int_SF (FF l2 f) l2 = Int_SF lf2 l2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a
Int_SF (FF l1 f) l1 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a
Int_SF (FF (cons_Rlist l2 l1) f) (cons_Rlist l2 l1) = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a

Int_SF (FF l1 f) l1 = Int_SF lf1 l1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a
Int_SF (FF (cons_Rlist l2 l1) f) (cons_Rlist l2 l1) = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a

Int_SF (FF (cons_Rlist l2 l1) f) (cons_Rlist l2 l1) = Int_SF lf3 l3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a

adapted_couple ?f ?a ?b (cons_Rlist l2 l1) (FF (cons_Rlist l2 l1) f)
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a
adapted_couple ?f ?a ?b l3 lf3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
Hnle:~ a <= b
Hnle':~ b <= c
Hnle'':~ a <= c
H:c < b
H0:b < a

adapted_couple f c a l3 lf3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3

(if Rle_dec a c then Int_SF lf3 l3 else - Int_SF lf3 l3) = RiemannInt_SF {| fe := f; pre := pr3 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
r:a <= c

Int_SF lf3 l3 = Int_SF (subdivision_val {| fe := f; pre := pr3 |}) (subdivision {| fe := f; pre := pr3 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
n:~ a <= c
- Int_SF lf3 l3 = - Int_SF (subdivision_val {| fe := f; pre := pr3 |}) (subdivision {| fe := f; pre := pr3 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
r:a <= c

adapted_couple ?f ?a ?b l3 lf3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
r:a <= c
adapted_couple ?f ?a ?b (subdivision {| fe := f; pre := pr3 |}) (subdivision_val {| fe := f; pre := pr3 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
n:~ a <= c
- Int_SF lf3 l3 = - Int_SF (subdivision_val {| fe := f; pre := pr3 |}) (subdivision {| fe := f; pre := pr3 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
r:a <= c

adapted_couple f a c (subdivision {| fe := f; pre := pr3 |}) (subdivision_val {| fe := f; pre := pr3 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
n:~ a <= c
- Int_SF lf3 l3 = - Int_SF (subdivision_val {| fe := f; pre := pr3 |}) (subdivision {| fe := f; pre := pr3 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
n:~ a <= c

- Int_SF lf3 l3 = - Int_SF (subdivision_val {| fe := f; pre := pr3 |}) (subdivision {| fe := f; pre := pr3 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
n:~ a <= c

adapted_couple ?f ?a ?b l3 lf3
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
n:~ a <= c
adapted_couple ?f ?a ?b (subdivision {| fe := f; pre := pr3 |}) (subdivision_val {| fe := f; pre := pr3 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
n:~ a <= c

adapted_couple f a c (subdivision {| fe := f; pre := pr3 |}) (subdivision_val {| fe := f; pre := pr3 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3

(if Rle_dec b c then Int_SF lf2 l2 else - Int_SF lf2 l2) = RiemannInt_SF {| fe := f; pre := pr2 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
r:b <= c

Int_SF lf2 l2 = Int_SF (subdivision_val {| fe := f; pre := pr2 |}) (subdivision {| fe := f; pre := pr2 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
n:~ b <= c
- Int_SF lf2 l2 = - Int_SF (subdivision_val {| fe := f; pre := pr2 |}) (subdivision {| fe := f; pre := pr2 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
r:b <= c

adapted_couple ?f ?a ?b l2 lf2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
r:b <= c
adapted_couple ?f ?a ?b (subdivision {| fe := f; pre := pr2 |}) (subdivision_val {| fe := f; pre := pr2 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
n:~ b <= c
- Int_SF lf2 l2 = - Int_SF (subdivision_val {| fe := f; pre := pr2 |}) (subdivision {| fe := f; pre := pr2 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
r:b <= c

adapted_couple f b c (subdivision {| fe := f; pre := pr2 |}) (subdivision_val {| fe := f; pre := pr2 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
n:~ b <= c
- Int_SF lf2 l2 = - Int_SF (subdivision_val {| fe := f; pre := pr2 |}) (subdivision {| fe := f; pre := pr2 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
n:~ b <= c

- Int_SF lf2 l2 = - Int_SF (subdivision_val {| fe := f; pre := pr2 |}) (subdivision {| fe := f; pre := pr2 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
n:~ b <= c

adapted_couple ?f ?a ?b l2 lf2
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
n:~ b <= c
adapted_couple ?f ?a ?b (subdivision {| fe := f; pre := pr2 |}) (subdivision_val {| fe := f; pre := pr2 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
n:~ b <= c

adapted_couple f b c (subdivision {| fe := f; pre := pr2 |}) (subdivision_val {| fe := f; pre := pr2 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3

(if Rle_dec a b then Int_SF lf1 l1 else - Int_SF lf1 l1) = RiemannInt_SF {| fe := f; pre := pr1 |}
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
r:a <= b

Int_SF lf1 l1 = Int_SF (subdivision_val {| fe := f; pre := pr1 |}) (subdivision {| fe := f; pre := pr1 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
n:~ a <= b
- Int_SF lf1 l1 = - Int_SF (subdivision_val {| fe := f; pre := pr1 |}) (subdivision {| fe := f; pre := pr1 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
r:a <= b

adapted_couple ?f ?a ?b l1 lf1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
r:a <= b
adapted_couple ?f ?a ?b (subdivision {| fe := f; pre := pr1 |}) (subdivision_val {| fe := f; pre := pr1 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
n:~ a <= b
- Int_SF lf1 l1 = - Int_SF (subdivision_val {| fe := f; pre := pr1 |}) (subdivision {| fe := f; pre := pr1 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
r:a <= b

adapted_couple f a b (subdivision {| fe := f; pre := pr1 |}) (subdivision_val {| fe := f; pre := pr1 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
n:~ a <= b
- Int_SF lf1 l1 = - Int_SF (subdivision_val {| fe := f; pre := pr1 |}) (subdivision {| fe := f; pre := pr1 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
n:~ a <= b

- Int_SF lf1 l1 = - Int_SF (subdivision_val {| fe := f; pre := pr1 |}) (subdivision {| fe := f; pre := pr1 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
n:~ a <= b

adapted_couple ?f ?a ?b l1 lf1
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
n:~ a <= b
adapted_couple ?f ?a ?b (subdivision {| fe := f; pre := pr1 |}) (subdivision_val {| fe := f; pre := pr1 |})
f:R -> R
a, b, c:R
pr1:IsStepFun f a b
pr2:IsStepFun f b c
pr3:IsStepFun f a c
l1, lf1:Rlist
H1:adapted_couple f a b l1 lf1
l2, lf2:Rlist
H2:adapted_couple f b c l2 lf2
l3, lf3:Rlist
H3:adapted_couple f a c l3 lf3
n:~ a <= b

adapted_couple f a b (subdivision {| fe := f; pre := pr1 |}) (subdivision_val {| fe := f; pre := pr1 |})
change (adapted_couple (mkStepFun pr1) a b (subdivision (mkStepFun pr1)) (subdivision_val (mkStepFun pr1))); apply StepFun_P1. Qed.

forall (f : R -> R) (a b c : R), IsStepFun f a b -> a <= c <= b -> IsStepFun f a c

forall (f : R -> R) (a b c : R), IsStepFun f a b -> a <= c <= b -> IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
H:a <= c <= b

a <= b
f:R -> R
a, b, c:R
X:IsStepFun f a b
H:a <= c <= b
H0:a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
H:a <= c <= b
H0:a <= b

IsStepFun f a c
f:R -> R
a, b, c:R
H0:a <= b
H:a <= c
H1:c <= b
l1, lf1:Rlist
H2:adapted_couple f a b l1 lf1

(forall (l2 lf2 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 l2 lf2 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}) -> IsStepFun f a c
f:R -> R
a, b, c:R
H0:a <= b
H:a <= c
H1:c <= b
l1, lf1:Rlist
H2:adapted_couple f a b l1 lf1
forall (l0 lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 l0 lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l2 : Rlist & adapted_couple f0 a0 c0 l l2}}
f:R -> R
a, b, c:R
H0:a <= b
H:a <= c
H1:c <= b
l1, lf1:Rlist
H2:adapted_couple f a b l1 lf1
X:forall (l0 lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 l0 lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l2 : Rlist & adapted_couple f0 a0 c0 l l2}}

adapted_couple f a ?b ?l1 ?lf1
f:R -> R
a, b, c:R
H0:a <= b
H:a <= c
H1:c <= b
l1, lf1:Rlist
H2:adapted_couple f a b l1 lf1
X:forall (l0 lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 l0 lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l2 : Rlist & adapted_couple f0 a0 c0 l l2}}
a <= c <= ?b
f:R -> R
a, b, c:R
H0:a <= b
H:a <= c
H1:c <= b
l1, lf1:Rlist
H2:adapted_couple f a b l1 lf1
forall (l0 lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 l0 lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l2 : Rlist & adapted_couple f0 a0 c0 l l2}}
f:R -> R
a, b, c:R
H0:a <= b
H:a <= c
H1:c <= b
l1, lf1:Rlist
H2:adapted_couple f a b l1 lf1
X:forall (l0 lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 l0 lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l2 : Rlist & adapted_couple f0 a0 c0 l l2}}

a <= c <= b
f:R -> R
a, b, c:R
H0:a <= b
H:a <= c
H1:c <= b
l1, lf1:Rlist
H2:adapted_couple f a b l1 lf1
forall (l0 lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 l0 lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l2 : Rlist & adapted_couple f0 a0 c0 l l2}}
f:R -> R
a, b, c:R
H0:a <= b
H:a <= c
H1:c <= b
l1, lf1:Rlist
H2:adapted_couple f a b l1 lf1

forall (l0 lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 l0 lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l2 : Rlist & adapted_couple f0 a0 c0 l l2}}
l1:Rlist

forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b nil lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
forall (r : R) (r0 : Rlist), (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r0 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r0) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist

forall (r : R) (r0 : Rlist), (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r0 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r0) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist

(forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b nil lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r nil) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b

a = b
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b
{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)

a <= b
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)
H7:a <= b
a = b
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b
{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)
H7:a <= b

a = b
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b
{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)
H7:a <= b

Rmin a b = b
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)
H7:a <= b
Rmin a b = a
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b
{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)
H7:a <= b

Rmin a b = Rmax a b
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)
H7:a <= b
Rmax a b = b
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)
H7:a <= b
Rmin a b = a
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b
{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)
H7:a <= b

Rmax a b = b
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)
H7:a <= b
Rmin a b = a
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b
{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)
H7:a <= b

Rmin a b = a
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b
{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b

{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b

c = b
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b
H2:c = b
adapted_couple f a c (cons r nil) lf1
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b
H2:c = b

adapted_couple f a c (cons r nil) lf1
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist

forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf1 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 (cons r1 r2) lf1 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) nil
H0:a <= c <= b

{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 (cons r1 r2) lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 (cons r1 r2) lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}

{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 (cons r1 r2) lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b

{c <= r1} + {r1 < c}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 (cons r1 r2) lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 (cons r1 r2) lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a0 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}

{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1

{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)

r = a
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H6:r = a
adapted_couple f a c (cons r (cons c nil)) (cons r3 nil)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H6:r = a

adapted_couple f a c (cons r (cons c nil)) (cons r3 nil)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H6:r = a
H:a <= c
H0:c <= b

ordered_Rlist (cons r (cons c nil))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H6:r = a
H:a <= c
H0:c <= b
pos_Rl (cons r (cons c nil)) 0 = Rmin a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H6:r = a
H:a <= c
H0:c <= b
pos_Rl (cons r (cons c nil)) (Init.Nat.pred (Rlength (cons r (cons c nil)))) = Rmax a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H6:r = a
H:a <= c
H0:c <= b
forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons c nil))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons c nil)) i) (pos_Rl (cons r (cons c nil)) (S i))) (pos_Rl (cons r3 nil) i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H6:r = a
H:a <= c
H0:c <= b

pos_Rl (cons r (cons c nil)) 0 = Rmin a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H6:r = a
H:a <= c
H0:c <= b
pos_Rl (cons r (cons c nil)) (Init.Nat.pred (Rlength (cons r (cons c nil)))) = Rmax a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H6:r = a
H:a <= c
H0:c <= b
forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons c nil))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons c nil)) i) (pos_Rl (cons r (cons c nil)) (S i))) (pos_Rl (cons r3 nil) i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H6:r = a
H:a <= c
H0:c <= b

pos_Rl (cons r (cons c nil)) (Init.Nat.pred (Rlength (cons r (cons c nil)))) = Rmax a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H6:r = a
H:a <= c
H0:c <= b
forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons c nil))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons c nil)) i) (pos_Rl (cons r (cons c nil)) (S i))) (pos_Rl (cons r3 nil) i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H6:r = a
H:a <= c
H0:c <= b

forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons c nil))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons c nil)) i) (pos_Rl (cons r (cons c nil)) (S i))) (pos_Rl (cons r3 nil) i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H6:r = a
H:a <= c
H0:c <= b
i:nat
H8:(i < 1)%nat
x:R
H9:pos_Rl (cons r (cons c nil)) i < x < pos_Rl (cons r (cons c nil)) (S i)
H11:i = 0%nat

f x = pos_Rl (cons r3 nil) 0
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H6:r = a
H:a <= c
H0:c <= b
i:nat
H8:(i < 1)%nat
x:R
H9:pos_Rl (cons r (cons c nil)) i < x < pos_Rl (cons r (cons c nil)) (S i)
m:nat
H11:(S i <= 0)%nat
H10:m = 0%nat
f x = pos_Rl (cons r3 nil) i
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H6:r = a
H:a <= c
H0:c <= b
i:nat
H8:(i < 1)%nat
x:R
H9:pos_Rl (cons r (cons c nil)) i < x < pos_Rl (cons r (cons c nil)) (S i)
H11:i = 0%nat
H10:(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) 0) (pos_Rl (cons r (cons r1 r2)) 1)) (pos_Rl (cons r3 lf1) 0)

(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H6:r = a
H:a <= c
H0:c <= b
i:nat
H8:(i < 1)%nat
x:R
H9:pos_Rl (cons r (cons c nil)) i < x < pos_Rl (cons r (cons c nil)) (S i)
H11:i = 0%nat
H10:(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) 0) (pos_Rl (cons r (cons r1 r2)) 1)) (pos_Rl (cons r3 lf1) 0)
H12:(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat
f x = r3
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H6:r = a
H:a <= c
H0:c <= b
i:nat
H8:(i < 1)%nat
x:R
H9:pos_Rl (cons r (cons c nil)) i < x < pos_Rl (cons r (cons c nil)) (S i)
m:nat
H11:(S i <= 0)%nat
H10:m = 0%nat
f x = pos_Rl (cons r3 nil) i
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H6:r = a
H:a <= c
H0:c <= b
i:nat
H8:(i < 1)%nat
x:R
H9:pos_Rl (cons r (cons c nil)) i < x < pos_Rl (cons r (cons c nil)) (S i)
H11:i = 0%nat
H10:(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) 0) (pos_Rl (cons r (cons r1 r2)) 1)) (pos_Rl (cons r3 lf1) 0)
H12:(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat

f x = r3
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H6:r = a
H:a <= c
H0:c <= b
i:nat
H8:(i < 1)%nat
x:R
H9:pos_Rl (cons r (cons c nil)) i < x < pos_Rl (cons r (cons c nil)) (S i)
m:nat
H11:(S i <= 0)%nat
H10:m = 0%nat
f x = pos_Rl (cons r3 nil) i
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H6:r = a
H:a <= c
H0:c <= b
i:nat
H8:(i < 1)%nat
x:R
H11:i = 0%nat
H10:(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) 0) (pos_Rl (cons r (cons r1 r2)) 1)) (pos_Rl (cons r3 lf1) 0)
H12:(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat
H9:r < x
H13:x < c

x < r1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H6:r = a
H:a <= c
H0:c <= b
i:nat
H8:(i < 1)%nat
x:R
H9:pos_Rl (cons r (cons c nil)) i < x < pos_Rl (cons r (cons c nil)) (S i)
m:nat
H11:(S i <= 0)%nat
H10:m = 0%nat
f x = pos_Rl (cons r3 nil) i
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H6:r = a
H:a <= c
H0:c <= b
i:nat
H8:(i < 1)%nat
x:R
H9:pos_Rl (cons r (cons c nil)) i < x < pos_Rl (cons r (cons c nil)) (S i)
m:nat
H11:(S i <= 0)%nat
H10:m = 0%nat

f x = pos_Rl (cons r3 nil) i
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c

{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c

adapted_couple f r1 b (cons r1 r2) lf1 -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c

r1 <= c <= b -> adapted_couple f r1 b (cons r1 r2) lf1 -> {l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1

{l : Rlist & {l0 : Rlist & adapted_couple f a c l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist l1'
H12:pos_Rl l1' 0 = Rmin r1 c
H11:pos_Rl l1' (Init.Nat.pred (Rlength l1')) = Rmax r1 c
H13:Rlength l1' = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength l1'))%nat -> constant_D_eq f (open_interval (pos_Rl l1' i) (pos_Rl l1' (S i))) (pos_Rl lf1' i)

a <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist l1'
H12:pos_Rl l1' 0 = Rmin r1 c
H11:pos_Rl l1' (Init.Nat.pred (Rlength l1')) = Rmax r1 c
H13:Rlength l1' = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength l1'))%nat -> constant_D_eq f (open_interval (pos_Rl l1' i) (pos_Rl l1' (S i))) (pos_Rl lf1' i)
H14:a <= b
adapted_couple f a c (cons r l1') (cons r3 lf1')
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist l1'
H12:pos_Rl l1' 0 = Rmin r1 c
H11:pos_Rl l1' (Init.Nat.pred (Rlength l1')) = Rmax r1 c
H13:Rlength l1' = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength l1'))%nat -> constant_D_eq f (open_interval (pos_Rl l1' i) (pos_Rl l1' (S i))) (pos_Rl lf1' i)
H14:a <= b

adapted_couple f a c (cons r l1') (cons r3 lf1')
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist l1'
H12:pos_Rl l1' 0 = Rmin r1 c
H11:pos_Rl l1' (Init.Nat.pred (Rlength l1')) = Rmax r1 c
H13:Rlength l1' = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength l1'))%nat -> constant_D_eq f (open_interval (pos_Rl l1' i) (pos_Rl l1' (S i))) (pos_Rl lf1' i)
H14:a <= b

r = a
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist l1'
H12:pos_Rl l1' 0 = Rmin r1 c
H11:pos_Rl l1' (Init.Nat.pred (Rlength l1')) = Rmax r1 c
H13:Rlength l1' = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength l1'))%nat -> constant_D_eq f (open_interval (pos_Rl l1' i) (pos_Rl l1' (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
adapted_couple f a c (cons r l1') (cons r3 lf1')
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist l1'
H12:pos_Rl l1' 0 = Rmin r1 c
H11:pos_Rl l1' (Init.Nat.pred (Rlength l1')) = Rmax r1 c
H13:Rlength l1' = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength l1'))%nat -> constant_D_eq f (open_interval (pos_Rl l1' i) (pos_Rl l1' (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a

adapted_couple f a c (cons r l1') (cons r3 lf1')
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist nil
H12:pos_Rl nil 0 = Rmin r1 c
H11:pos_Rl nil (Init.Nat.pred (Rlength nil)) = Rmax r1 c
H13:Rlength nil = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength nil))%nat -> constant_D_eq f (open_interval (pos_Rl nil i) (pos_Rl nil (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a

adapted_couple f a c (cons r nil) (cons r3 lf1')
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
Hrecl1':ordered_Rlist l1' -> pos_Rl l1' 0 = Rmin r1 c -> pos_Rl l1' (Init.Nat.pred (Rlength l1')) = Rmax r1 c -> Rlength l1' = S (Rlength lf1') -> (forall i : nat, (i < Init.Nat.pred (Rlength l1'))%nat -> constant_D_eq f (open_interval (pos_Rl l1' i) (pos_Rl l1' (S i))) (pos_Rl lf1' i)) -> adapted_couple f a c (cons r l1') (cons r3 lf1')
adapted_couple f a c (cons r (cons r4 l1')) (cons r3 lf1')
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
Hrecl1':ordered_Rlist l1' -> pos_Rl l1' 0 = Rmin r1 c -> pos_Rl l1' (Init.Nat.pred (Rlength l1')) = Rmax r1 c -> Rlength l1' = S (Rlength lf1') -> (forall i : nat, (i < Init.Nat.pred (Rlength l1'))%nat -> constant_D_eq f (open_interval (pos_Rl l1' i) (pos_Rl l1' (S i))) (pos_Rl lf1' i)) -> adapted_couple f a c (cons r l1') (cons r3 lf1')

adapted_couple f a c (cons r (cons r4 l1')) (cons r3 lf1')
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a

ordered_Rlist (cons r (cons r4 l1'))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
pos_Rl (cons r (cons r4 l1')) 0 = Rmin a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
pos_Rl (cons r (cons r4 l1')) (Init.Nat.pred (Rlength (cons r (cons r4 l1')))) = Rmax a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
Rlength (cons r (cons r4 l1')) = S (Rlength (cons r3 lf1'))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r4 l1'))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r4 l1')) i) (pos_Rl (cons r (cons r4 l1')) (S i))) (pos_Rl (cons r3 lf1') i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
H:(0 < S (Rlength l1'))%nat

pos_Rl (cons r (cons r4 l1')) 0 <= pos_Rl (cons r (cons r4 l1')) 1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i0) (pos_Rl (cons r4 l1') (S i0))) (pos_Rl lf1' i0)
H14:a <= b
H16:r = a
i:nat
H:(S i < S (Rlength l1'))%nat
Hreci:(i < S (Rlength l1'))%nat -> pos_Rl (cons r (cons r4 l1')) i <= pos_Rl (cons r (cons r4 l1')) (S i)
pos_Rl (cons r (cons r4 l1')) (S i) <= pos_Rl (cons r (cons r4 l1')) (S (S i))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
pos_Rl (cons r (cons r4 l1')) 0 = Rmin a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
pos_Rl (cons r (cons r4 l1')) (Init.Nat.pred (Rlength (cons r (cons r4 l1')))) = Rmax a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
Rlength (cons r (cons r4 l1')) = S (Rlength (cons r3 lf1'))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r4 l1'))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r4 l1')) i) (pos_Rl (cons r (cons r4 l1')) (S i))) (pos_Rl (cons r3 lf1') i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
H:(0 < S (Rlength l1'))%nat

r <= r1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
H:(0 < S (Rlength l1'))%nat
r1 = r4
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i0) (pos_Rl (cons r4 l1') (S i0))) (pos_Rl lf1' i0)
H14:a <= b
H16:r = a
i:nat
H:(S i < S (Rlength l1'))%nat
Hreci:(i < S (Rlength l1'))%nat -> pos_Rl (cons r (cons r4 l1')) i <= pos_Rl (cons r (cons r4 l1')) (S i)
pos_Rl (cons r (cons r4 l1')) (S i) <= pos_Rl (cons r (cons r4 l1')) (S (S i))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
pos_Rl (cons r (cons r4 l1')) 0 = Rmin a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
pos_Rl (cons r (cons r4 l1')) (Init.Nat.pred (Rlength (cons r (cons r4 l1')))) = Rmax a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
Rlength (cons r (cons r4 l1')) = S (Rlength (cons r3 lf1'))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r4 l1'))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r4 l1')) i) (pos_Rl (cons r (cons r4 l1')) (S i))) (pos_Rl (cons r3 lf1') i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
H:(0 < S (Rlength l1'))%nat

(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
H:(0 < S (Rlength l1'))%nat
r1 = r4
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i0) (pos_Rl (cons r4 l1') (S i0))) (pos_Rl lf1' i0)
H14:a <= b
H16:r = a
i:nat
H:(S i < S (Rlength l1'))%nat
Hreci:(i < S (Rlength l1'))%nat -> pos_Rl (cons r (cons r4 l1')) i <= pos_Rl (cons r (cons r4 l1')) (S i)
pos_Rl (cons r (cons r4 l1')) (S i) <= pos_Rl (cons r (cons r4 l1')) (S (S i))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
pos_Rl (cons r (cons r4 l1')) 0 = Rmin a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
pos_Rl (cons r (cons r4 l1')) (Init.Nat.pred (Rlength (cons r (cons r4 l1')))) = Rmax a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
Rlength (cons r (cons r4 l1')) = S (Rlength (cons r3 lf1'))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r4 l1'))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r4 l1')) i) (pos_Rl (cons r (cons r4 l1')) (S i))) (pos_Rl (cons r3 lf1') i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
H:(0 < S (Rlength l1'))%nat

r1 = r4
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i0) (pos_Rl (cons r4 l1') (S i0))) (pos_Rl lf1' i0)
H14:a <= b
H16:r = a
i:nat
H:(S i < S (Rlength l1'))%nat
Hreci:(i < S (Rlength l1'))%nat -> pos_Rl (cons r (cons r4 l1')) i <= pos_Rl (cons r (cons r4 l1')) (S i)
pos_Rl (cons r (cons r4 l1')) (S i) <= pos_Rl (cons r (cons r4 l1')) (S (S i))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
pos_Rl (cons r (cons r4 l1')) 0 = Rmin a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
pos_Rl (cons r (cons r4 l1')) (Init.Nat.pred (Rlength (cons r (cons r4 l1')))) = Rmax a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
Rlength (cons r (cons r4 l1')) = S (Rlength (cons r3 lf1'))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r4 l1'))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r4 l1')) i) (pos_Rl (cons r (cons r4 l1')) (S i))) (pos_Rl (cons r3 lf1') i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i0) (pos_Rl (cons r4 l1') (S i0))) (pos_Rl lf1' i0)
H14:a <= b
H16:r = a
i:nat
H:(S i < S (Rlength l1'))%nat
Hreci:(i < S (Rlength l1'))%nat -> pos_Rl (cons r (cons r4 l1')) i <= pos_Rl (cons r (cons r4 l1')) (S i)

pos_Rl (cons r (cons r4 l1')) (S i) <= pos_Rl (cons r (cons r4 l1')) (S (S i))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
pos_Rl (cons r (cons r4 l1')) 0 = Rmin a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
pos_Rl (cons r (cons r4 l1')) (Init.Nat.pred (Rlength (cons r (cons r4 l1')))) = Rmax a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
Rlength (cons r (cons r4 l1')) = S (Rlength (cons r3 lf1'))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r4 l1'))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r4 l1')) i) (pos_Rl (cons r (cons r4 l1')) (S i))) (pos_Rl (cons r3 lf1') i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a

pos_Rl (cons r (cons r4 l1')) 0 = Rmin a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
pos_Rl (cons r (cons r4 l1')) (Init.Nat.pred (Rlength (cons r (cons r4 l1')))) = Rmax a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
Rlength (cons r (cons r4 l1')) = S (Rlength (cons r3 lf1'))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r4 l1'))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r4 l1')) i) (pos_Rl (cons r (cons r4 l1')) (S i))) (pos_Rl (cons r3 lf1') i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a

pos_Rl (cons r (cons r4 l1')) (Init.Nat.pred (Rlength (cons r (cons r4 l1')))) = Rmax a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
Rlength (cons r (cons r4 l1')) = S (Rlength (cons r3 lf1'))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r4 l1'))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r4 l1')) i) (pos_Rl (cons r (cons r4 l1')) (S i))) (pos_Rl (cons r3 lf1') i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a

pos_Rl (cons r (cons r4 l1')) (Init.Nat.pred (Rlength (cons r (cons r4 l1')))) = Rmax r1 c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
Rmax r1 c = Rmax a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
Rlength (cons r (cons r4 l1')) = S (Rlength (cons r3 lf1'))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r4 l1'))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r4 l1')) i) (pos_Rl (cons r (cons r4 l1')) (S i))) (pos_Rl (cons r3 lf1') i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a

Rmax r1 c = Rmax a c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
Rlength (cons r (cons r4 l1')) = S (Rlength (cons r3 lf1'))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r4 l1'))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r4 l1')) i) (pos_Rl (cons r (cons r4 l1')) (S i))) (pos_Rl (cons r3 lf1') i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a

Rlength (cons r (cons r4 l1')) = S (Rlength (cons r3 lf1'))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r4 l1'))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r4 l1')) i) (pos_Rl (cons r (cons r4 l1')) (S i))) (pos_Rl (cons r3 lf1') i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a

forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r4 l1'))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r4 l1')) i) (pos_Rl (cons r (cons r4 l1')) (S i))) (pos_Rl (cons r3 lf1') i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
H:(0 < S (Rlength l1'))%nat
x:R
H4:pos_Rl (cons r (cons r4 l1')) 0 < x < pos_Rl (cons r (cons r4 l1')) 1

f x = pos_Rl (cons r3 lf1') 0
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i0) (pos_Rl (cons r4 l1') (S i0))) (pos_Rl lf1' i0)
H14:a <= b
H16:r = a
i:nat
H:(S i < S (Rlength l1'))%nat
x:R
H4:pos_Rl (cons r (cons r4 l1')) (S i) < x < pos_Rl (cons r (cons r4 l1')) (S (S i))
Hreci:(i < S (Rlength l1'))%nat -> pos_Rl (cons r (cons r4 l1')) i < x < pos_Rl (cons r (cons r4 l1')) (S i) -> f x = pos_Rl (cons r3 lf1') i
f x = pos_Rl (cons r3 lf1') (S i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
H:(0 < S (Rlength l1'))%nat
x:R
H4:pos_Rl (cons r (cons r4 l1')) 0 < x < pos_Rl (cons r (cons r4 l1')) 1
H17:(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) 0) (pos_Rl (cons r (cons r1 r2)) 1)) (pos_Rl (cons r3 lf1) 0)

(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
H:(0 < S (Rlength l1'))%nat
x:R
H4:pos_Rl (cons r (cons r4 l1')) 0 < x < pos_Rl (cons r (cons r4 l1')) 1
H17:(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) 0) (pos_Rl (cons r (cons r1 r2)) 1)) (pos_Rl (cons r3 lf1) 0)
H18:(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat
f x = r3
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i0) (pos_Rl (cons r4 l1') (S i0))) (pos_Rl lf1' i0)
H14:a <= b
H16:r = a
i:nat
H:(S i < S (Rlength l1'))%nat
x:R
H4:pos_Rl (cons r (cons r4 l1')) (S i) < x < pos_Rl (cons r (cons r4 l1')) (S (S i))
Hreci:(i < S (Rlength l1'))%nat -> pos_Rl (cons r (cons r4 l1')) i < x < pos_Rl (cons r (cons r4 l1')) (S i) -> f x = pos_Rl (cons r3 lf1') i
f x = pos_Rl (cons r3 lf1') (S i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
H:(0 < S (Rlength l1'))%nat
x:R
H4:pos_Rl (cons r (cons r4 l1')) 0 < x < pos_Rl (cons r (cons r4 l1')) 1
H17:(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) 0) (pos_Rl (cons r (cons r1 r2)) 1)) (pos_Rl (cons r3 lf1) 0)
H18:(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat

f x = r3
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i0) (pos_Rl (cons r4 l1') (S i0))) (pos_Rl lf1' i0)
H14:a <= b
H16:r = a
i:nat
H:(S i < S (Rlength l1'))%nat
x:R
H4:pos_Rl (cons r (cons r4 l1')) (S i) < x < pos_Rl (cons r (cons r4 l1')) (S (S i))
Hreci:(i < S (Rlength l1'))%nat -> pos_Rl (cons r (cons r4 l1')) i < x < pos_Rl (cons r (cons r4 l1')) (S i) -> f x = pos_Rl (cons r3 lf1') i
f x = pos_Rl (cons r3 lf1') (S i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
H:(0 < S (Rlength l1'))%nat
x:R
H17:(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) 0) (pos_Rl (cons r (cons r1 r2)) 1)) (pos_Rl (cons r3 lf1) 0)
H18:(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat
H4:r < x
H19:x < r4

x < r4
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
H:(0 < S (Rlength l1'))%nat
x:R
H17:(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) 0) (pos_Rl (cons r (cons r1 r2)) 1)) (pos_Rl (cons r3 lf1) 0)
H18:(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat
H4:r < x
H19:x < r4
r4 = r1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i0) (pos_Rl (cons r4 l1') (S i0))) (pos_Rl lf1' i0)
H14:a <= b
H16:r = a
i:nat
H:(S i < S (Rlength l1'))%nat
x:R
H4:pos_Rl (cons r (cons r4 l1')) (S i) < x < pos_Rl (cons r (cons r4 l1')) (S (S i))
Hreci:(i < S (Rlength l1'))%nat -> pos_Rl (cons r (cons r4 l1')) i < x < pos_Rl (cons r (cons r4 l1')) (S i) -> f x = pos_Rl (cons r3 lf1') i
f x = pos_Rl (cons r3 lf1') (S i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i : nat, (i < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i))) (pos_Rl lf1' i)
H14:a <= b
H16:r = a
H:(0 < S (Rlength l1'))%nat
x:R
H17:(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) 0) (pos_Rl (cons r (cons r1 r2)) 1)) (pos_Rl (cons r3 lf1) 0)
H18:(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat
H4:r < x
H19:x < r4

r4 = r1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i0) (pos_Rl (cons r4 l1') (S i0))) (pos_Rl lf1' i0)
H14:a <= b
H16:r = a
i:nat
H:(S i < S (Rlength l1'))%nat
x:R
H4:pos_Rl (cons r (cons r4 l1')) (S i) < x < pos_Rl (cons r (cons r4 l1')) (S (S i))
Hreci:(i < S (Rlength l1'))%nat -> pos_Rl (cons r (cons r4 l1')) i < x < pos_Rl (cons r (cons r4 l1')) (S i) -> f x = pos_Rl (cons r3 lf1') i
f x = pos_Rl (cons r3 lf1') (S i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i0) (pos_Rl (cons r4 l1') (S i0))) (pos_Rl lf1' i0)
H14:a <= b
H16:r = a
i:nat
H:(S i < S (Rlength l1'))%nat
x:R
H4:pos_Rl (cons r (cons r4 l1')) (S i) < x < pos_Rl (cons r (cons r4 l1')) (S (S i))
Hreci:(i < S (Rlength l1'))%nat -> pos_Rl (cons r (cons r4 l1')) i < x < pos_Rl (cons r (cons r4 l1')) (S i) -> f x = pos_Rl (cons r3 lf1') i

f x = pos_Rl (cons r3 lf1') (S i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i0) (pos_Rl (cons r4 l1') (S i0))) (pos_Rl lf1' i0)
H14:a <= b
H16:r = a
i:nat
H:(S i < S (Rlength l1'))%nat
x:R
H4:pos_Rl (cons r (cons r4 l1')) (S i) < x < pos_Rl (cons r (cons r4 l1')) (S (S i))

(i < Init.Nat.pred (Rlength (cons r4 l1')))%nat
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i0) (pos_Rl (cons r4 l1') (S i0))) (pos_Rl lf1' i0)
H14:a <= b
H16:r = a
i:nat
H:(S i < S (Rlength l1'))%nat
x:R
H4:pos_Rl (cons r (cons r4 l1')) (S i) < x < pos_Rl (cons r (cons r4 l1')) (S (S i))
open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i)) x
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
H2:r1 <= c <= b
H3:adapted_couple f r1 b (cons r1 r2) lf1
r4:R
l1', lf1':Rlist
H5:ordered_Rlist (cons r (cons r1 r2))
H7:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H6:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H8:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H10:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
H9:ordered_Rlist (cons r4 l1')
H12:pos_Rl (cons r4 l1') 0 = Rmin r1 c
H11:pos_Rl (cons r4 l1') (Init.Nat.pred (Rlength (cons r4 l1'))) = Rmax r1 c
H13:Rlength (cons r4 l1') = S (Rlength lf1')
H15:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r4 l1')))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r4 l1') i0) (pos_Rl (cons r4 l1') (S i0))) (pos_Rl lf1' i0)
H14:a <= b
H16:r = a
i:nat
H:(S i < S (Rlength l1'))%nat
x:R
H4:pos_Rl (cons r (cons r4 l1')) (S i) < x < pos_Rl (cons r (cons r4 l1')) (S (S i))

open_interval (pos_Rl (cons r4 l1') i) (pos_Rl (cons r4 l1') (S i)) x
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c

r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c

r1 <= c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c

c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 a1 c0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c

adapted_couple f r1 b (cons r1 r2) lf1
eapply StepFun_P7; [ elim H0; intros; apply Rle_trans with c; [ apply H2 | apply H3 ] | apply H ]. Qed.

forall (f : R -> R) (a b c : R), IsStepFun f a b -> a <= c <= b -> IsStepFun f c b

forall (f : R -> R) (a b c : R), IsStepFun f a b -> a <= c <= b -> IsStepFun f c b
f:R -> R
a, b, c:R
X:IsStepFun f a b
H:a <= c <= b

a <= b
f:R -> R
a, b, c:R
X:IsStepFun f a b
H:a <= c <= b
H0:a <= b
IsStepFun f c b
f:R -> R
a, b, c:R
X:IsStepFun f a b
H:a <= c <= b
H0:a <= b

IsStepFun f c b
f:R -> R
a, b, c:R
H0:a <= b
H:a <= c
H1:c <= b
l1, lf1:Rlist
H2:adapted_couple f a b l1 lf1

(forall (l2 lf2 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 l2 lf2 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}) -> IsStepFun f c b
f:R -> R
a, b, c:R
H0:a <= b
H:a <= c
H1:c <= b
l1, lf1:Rlist
H2:adapted_couple f a b l1 lf1
forall (l0 lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 l0 lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l2 : Rlist & adapted_couple f0 c0 b0 l l2}}
f:R -> R
a, b, c:R
H0:a <= b
H:a <= c
H1:c <= b
l1, lf1:Rlist
H2:adapted_couple f a b l1 lf1

forall (l0 lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 l0 lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l2 : Rlist & adapted_couple f0 c0 b0 l l2}}
l1:Rlist

forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b nil lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
forall (r : R) (r0 : Rlist), (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r0 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r0) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist

forall (r : R) (r0 : Rlist), (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r0 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r0) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist

(forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b nil lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r nil) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b

a = b
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)

a <= b
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)
H7:a <= b
a = b
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)
H7:a <= b

a = b
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)
H7:a <= b

Rmin a b = b
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)
H7:a <= b
Rmin a b = a
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)
H7:a <= b

Rmin a b = Rmax a b
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)
H7:a <= b
Rmax a b = b
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)
H7:a <= b
Rmin a b = a
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)
H7:a <= b

Rmax a b = b
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)
H7:a <= b
Rmin a b = a
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:ordered_Rlist (cons r nil)
H3:r = Rmin a b
H2:r = Rmax a b
H4:Rlength (cons r nil) = S (Rlength lf1)
H6:forall i : nat, (i < Init.Nat.pred (Rlength (cons r nil)))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r nil) i) (pos_Rl (cons r nil) (S i))) (pos_Rl lf1 i)
H7:a <= b

Rmin a b = a
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b

{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b

c = b
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b
H2:c = b
adapted_couple f c b (cons r nil) lf1
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
X:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 nil lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r nil) lf1
H0:a <= c <= b
H1:a = b
H2:c = b

adapted_couple f c b (cons r nil) lf1
l1:Rlist
r:R
r0:Rlist
forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist

forall (r1 : R) (r2 : Rlist), ((forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b r2 lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> (forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r1 r2) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}) -> forall (lf1 : Rlist) (a b c : R) (f : R -> R), adapted_couple f a b (cons r (cons r1 r2)) lf1 -> a <= c <= b -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf1 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 (cons r1 r2) lf1 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) nil
H0:a <= c <= b

{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 (cons r1 r2) lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 (cons r1 r2) lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
Hreclf1:adapted_couple f a b (cons r (cons r1 r2)) lf1 -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}

{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 (cons r1 r2) lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b

{c <= r1} + {r1 < c}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 (cons r1 r2) lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a0 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a0 b0 (cons r1 r2) lf0 -> a0 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}

{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1

{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)

ordered_Rlist (cons c (cons r1 r2))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
pos_Rl (cons c (cons r1 r2)) 0 = Rmin c b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
pos_Rl (cons c (cons r1 r2)) (Init.Nat.pred (Rlength (cons c (cons r1 r2)))) = Rmax c b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
Rlength (cons c (cons r1 r2)) = S (Rlength (cons r3 lf1))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons c (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) i) (pos_Rl (cons c (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H:(0 < S (Rlength r2))%nat

pos_Rl (cons c (cons r1 r2)) 0 <= pos_Rl (cons c (cons r1 r2)) 1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
i:nat
H:(S i < S (Rlength r2))%nat
Hreci:(i < S (Rlength r2))%nat -> pos_Rl (cons c (cons r1 r2)) i <= pos_Rl (cons c (cons r1 r2)) (S i)
pos_Rl (cons c (cons r1 r2)) (S i) <= pos_Rl (cons c (cons r1 r2)) (S (S i))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
pos_Rl (cons c (cons r1 r2)) 0 = Rmin c b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
pos_Rl (cons c (cons r1 r2)) (Init.Nat.pred (Rlength (cons c (cons r1 r2)))) = Rmax c b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
Rlength (cons c (cons r1 r2)) = S (Rlength (cons r3 lf1))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons c (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) i) (pos_Rl (cons c (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
i:nat
H:(S i < S (Rlength r2))%nat
Hreci:(i < S (Rlength r2))%nat -> pos_Rl (cons c (cons r1 r2)) i <= pos_Rl (cons c (cons r1 r2)) (S i)

pos_Rl (cons c (cons r1 r2)) (S i) <= pos_Rl (cons c (cons r1 r2)) (S (S i))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
pos_Rl (cons c (cons r1 r2)) 0 = Rmin c b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
pos_Rl (cons c (cons r1 r2)) (Init.Nat.pred (Rlength (cons c (cons r1 r2)))) = Rmax c b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
Rlength (cons c (cons r1 r2)) = S (Rlength (cons r3 lf1))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons c (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) i) (pos_Rl (cons c (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)

pos_Rl (cons c (cons r1 r2)) 0 = Rmin c b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
pos_Rl (cons c (cons r1 r2)) (Init.Nat.pred (Rlength (cons c (cons r1 r2)))) = Rmax c b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
Rlength (cons c (cons r1 r2)) = S (Rlength (cons r3 lf1))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons c (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) i) (pos_Rl (cons c (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)

pos_Rl (cons c (cons r1 r2)) (Init.Nat.pred (Rlength (cons c (cons r1 r2)))) = Rmax c b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
Rlength (cons c (cons r1 r2)) = S (Rlength (cons r3 lf1))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons c (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) i) (pos_Rl (cons c (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)

pos_Rl (cons c (cons r1 r2)) (Init.Nat.pred (Rlength (cons c (cons r1 r2)))) = Rmax a b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
Rmax a b = Rmax c b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
Rlength (cons c (cons r1 r2)) = S (Rlength (cons r3 lf1))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons c (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) i) (pos_Rl (cons c (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)

Rmax a b = Rmax c b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
Rlength (cons c (cons r1 r2)) = S (Rlength (cons r3 lf1))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons c (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) i) (pos_Rl (cons c (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)

Rlength (cons c (cons r1 r2)) = S (Rlength (cons r3 lf1))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
forall i : nat, (i < Init.Nat.pred (Rlength (cons c (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) i) (pos_Rl (cons c (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)

forall i : nat, (i < Init.Nat.pred (Rlength (cons c (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) i) (pos_Rl (cons c (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H:(0 < S (Rlength r2))%nat

constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) 0) (pos_Rl (cons c (cons r1 r2)) 1)) (pos_Rl (cons r3 lf1) 0)
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
i:nat
H:(S i < S (Rlength r2))%nat
Hreci:(i < S (Rlength r2))%nat -> constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) i) (pos_Rl (cons c (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) (S i)) (pos_Rl (cons c (cons r1 r2)) (S (S i)))) (pos_Rl (cons r3 lf1) (S i))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H:(0 < S (Rlength r2))%nat
x:R
H6:pos_Rl (cons c (cons r1 r2)) 0 < x < pos_Rl (cons c (cons r1 r2)) 1

(0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H:(0 < S (Rlength r2))%nat
x:R
H6:pos_Rl (cons c (cons r1 r2)) 0 < x < pos_Rl (cons c (cons r1 r2)) 1
open_interval (pos_Rl (cons r (cons r1 r2)) 0) (pos_Rl (cons r (cons r1 r2)) 1) x
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
i:nat
H:(S i < S (Rlength r2))%nat
Hreci:(i < S (Rlength r2))%nat -> constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) i) (pos_Rl (cons c (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) (S i)) (pos_Rl (cons c (cons r1 r2)) (S (S i)))) (pos_Rl (cons r3 lf1) (S i))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H:(0 < S (Rlength r2))%nat
x:R
H6:pos_Rl (cons c (cons r1 r2)) 0 < x < pos_Rl (cons c (cons r1 r2)) 1

open_interval (pos_Rl (cons r (cons r1 r2)) 0) (pos_Rl (cons r (cons r1 r2)) 1) x
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
i:nat
H:(S i < S (Rlength r2))%nat
Hreci:(i < S (Rlength r2))%nat -> constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) i) (pos_Rl (cons c (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) (S i)) (pos_Rl (cons c (cons r1 r2)) (S (S i)))) (pos_Rl (cons r3 lf1) (S i))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H:(0 < S (Rlength r2))%nat
x:R
H6:c < x
H8:x < r1

a <= c
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H:(0 < S (Rlength r2))%nat
x:R
H6:c < x
H8:x < r1
a = r
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
i:nat
H:(S i < S (Rlength r2))%nat
Hreci:(i < S (Rlength r2))%nat -> constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) i) (pos_Rl (cons c (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) (S i)) (pos_Rl (cons c (cons r1 r2)) (S (S i)))) (pos_Rl (cons r3 lf1) (S i))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i : nat, (i < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i) (pos_Rl (cons r (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
H:(0 < S (Rlength r2))%nat
x:R
H6:c < x
H8:x < r1

a = r
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
i:nat
H:(S i < S (Rlength r2))%nat
Hreci:(i < S (Rlength r2))%nat -> constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) i) (pos_Rl (cons c (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)
constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) (S i)) (pos_Rl (cons c (cons r1 r2)) (S (S i)))) (pos_Rl (cons r3 lf1) (S i))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:c <= r1
H2:ordered_Rlist (cons r (cons r1 r2))
H4:pos_Rl (cons r (cons r1 r2)) 0 = Rmin a b
H3:pos_Rl (cons r (cons r1 r2)) (Init.Nat.pred (Rlength (cons r (cons r1 r2)))) = Rmax a b
H5:Rlength (cons r (cons r1 r2)) = S (Rlength (cons r3 lf1))
H7:forall i0 : nat, (i0 < Init.Nat.pred (Rlength (cons r (cons r1 r2))))%nat -> constant_D_eq f (open_interval (pos_Rl (cons r (cons r1 r2)) i0) (pos_Rl (cons r (cons r1 r2)) (S i0))) (pos_Rl (cons r3 lf1) i0)
i:nat
H:(S i < S (Rlength r2))%nat
Hreci:(i < S (Rlength r2))%nat -> constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) i) (pos_Rl (cons c (cons r1 r2)) (S i))) (pos_Rl (cons r3 lf1) i)

constant_D_eq f (open_interval (pos_Rl (cons c (cons r1 r2)) (S i)) (pos_Rl (cons c (cons r1 r2)) (S (S i)))) (pos_Rl (cons r3 lf1) (S i))
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c

{l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c

adapted_couple f r1 b (cons r1 r2) lf1 -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c

r1 <= c <= b -> adapted_couple f r1 b (cons r1 r2) lf1 -> {l : Rlist & {l0 : Rlist & adapted_couple f c b l l0}}
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c

r1 <= c <= b
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c
adapted_couple f r1 b (cons r1 r2) lf1
l1:Rlist
r:R
r0:Rlist
r1:R
r2:Rlist
X0:forall (lf0 : Rlist) (a1 b0 c0 : R) (f0 : R -> R), adapted_couple f0 a1 b0 (cons r1 r2) lf0 -> a1 <= c0 <= b0 -> {l : Rlist & {l0 : Rlist & adapted_couple f0 c0 b0 l l0}}
r3:R
lf1:Rlist
a, b, c:R
f:R -> R
H:adapted_couple f a b (cons r (cons r1 r2)) (cons r3 lf1)
H0:a <= c <= b
H1:{c <= r1} + {r1 < c}
a0:r1 < c

adapted_couple f r1 b (cons r1 r2) lf1
eapply StepFun_P7; [ elim H0; intros; apply Rle_trans with c; [ apply H2 | apply H3 ] | apply H ]. Qed.

forall (f : R -> R) (a b c : R), IsStepFun f a b -> IsStepFun f b c -> IsStepFun f a c

forall (f : R -> R) (a b c : R), IsStepFun f a b -> IsStepFun f b c -> IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
r:b <= c
r0:a <= b

IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
r:a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
r:b <= c
n:~ a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
n0:~ a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
r:a <= b

IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
r:b <= c
n:~ a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
n0:~ a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
r:a <= b
r0:a <= c

IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
r:a <= b
n0:~ a <= c
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
r:b <= c
n:~ a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
n0:~ a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
r:a <= b
r0:a <= c

a <= c <= b
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
r:a <= b
n0:~ a <= c
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
r:b <= c
n:~ a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
n0:~ a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
r:a <= b
n0:~ a <= c

IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
r:b <= c
n:~ a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
n0:~ a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
r:a <= b
n0:~ a <= c

IsStepFun f c b
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
r:a <= b
n0:~ a <= c
c <= a <= b
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
r:b <= c
n:~ a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
n0:~ a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
r:a <= b
n0:~ a <= c

c <= a <= b
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
r:b <= c
n:~ a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
n0:~ a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
r:b <= c
n:~ a <= b

IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
n0:~ a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
r:b <= c
n:~ a <= b
r0:a <= c

IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
r:b <= c
n:~ a <= b
n0:~ a <= c
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
n0:~ a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
r:b <= c
n:~ a <= b
r0:a <= c

b <= a <= c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
r:b <= c
n:~ a <= b
n0:~ a <= c
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
n0:~ a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
r:b <= c
n:~ a <= b
n0:~ a <= c

IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
n0:~ a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
r:b <= c
n:~ a <= b
n0:~ a <= c

IsStepFun f b a
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
r:b <= c
n:~ a <= b
n0:~ a <= c
b <= c <= a
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
n0:~ a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
r:b <= c
n:~ a <= b
n0:~ a <= c

b <= c <= a
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
n0:~ a <= b
IsStepFun f a c
f:R -> R
a, b, c:R
X:IsStepFun f a b
X0:IsStepFun f b c
n:~ b <= c
n0:~ a <= b

IsStepFun f a c
apply StepFun_P6; apply StepFun_P41 with b; auto with real || apply StepFun_P6; assumption. Qed.