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 Sumbool.
Require Import Rbase.
Require Import Rfunctions.
Require Import SeqSeries.
Require Import Ranalysis1.
Local Open Scope R_scope.

Fixpoint Dichotomy_lb (x y:R) (P:R -> bool) (N:nat) {struct N} : R :=
  match N with
    | O => x
    | S n =>
      let down := Dichotomy_lb x y P n in
        let up := Dichotomy_ub x y P n in
          let z := (down + up) / 2 in if P z then down else z
  end

  with Dichotomy_ub (x y:R) (P:R -> bool) (N:nat) {struct N} : R :=
    match N with
      | O => y
      | S n =>
        let down := Dichotomy_lb x y P n in
          let up := Dichotomy_ub x y P n in
            let z := (down + up) / 2 in if P z then z else up
    end.

Definition dicho_lb (x y:R) (P:R -> bool) (N:nat) : R := Dichotomy_lb x y P N.
Definition dicho_up (x y:R) (P:R -> bool) (N:nat) : R := Dichotomy_ub x y P N.

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

forall (x y : R) (P : R -> bool) (n : nat), x <= y -> dicho_lb x y P n <= dicho_up x y P n

forall (x y : R) (P : R -> bool) (n : nat), x <= y -> dicho_lb x y P n <= dicho_up x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y

dicho_lb x y P n <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y

dicho_lb x y P 0 <= dicho_up x y P 0
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n
dicho_lb x y P (S n) <= dicho_up x y P (S n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n

dicho_lb x y P (S n) <= dicho_up x y P (S n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n

(if P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) then Dichotomy_lb x y P n else (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) <= (if P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) then (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 else Dichotomy_ub x y P n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n

Dichotomy_lb x y P n <= (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n
(Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n

0 < 2
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n
2 * Dichotomy_lb x y P n <= 2 * ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n
(Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n

2 * Dichotomy_lb x y P n <= 2 * ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n
(Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n

2 * Dichotomy_lb x y P n <= (Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2 * 2
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n
(Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n

2 * Dichotomy_lb x y P n <= (Dichotomy_lb x y P n + Dichotomy_ub x y P n) * 1
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n
(Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n

2 * Dichotomy_lb x y P n <= Dichotomy_lb x y P n + Dichotomy_ub x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n
(Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n

Dichotomy_lb x y P n + Dichotomy_lb x y P n <= Dichotomy_lb x y P n + Dichotomy_ub x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n
(Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n

Dichotomy_lb x y P n <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n
(Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n

(Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n

0 < 2
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n
2 * ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2) <= 2 * Dichotomy_ub x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n

2 * ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2) <= 2 * Dichotomy_ub x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n

(Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2 * 2 <= 2 * Dichotomy_ub x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n

(Dichotomy_lb x y P n + Dichotomy_ub x y P n) * 1 <= 2 * Dichotomy_ub x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n

Dichotomy_lb x y P n + Dichotomy_ub x y P n <= 2 * Dichotomy_ub x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n

Dichotomy_lb x y P n + Dichotomy_ub x y P n <= Dichotomy_ub x y P n + Dichotomy_ub x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n

Dichotomy_ub x y P n + Dichotomy_lb x y P n <= Dichotomy_ub x y P n + Dichotomy_ub x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_lb x y P n <= dicho_up x y P n

Dichotomy_lb x y P n <= Dichotomy_ub x y P n
assumption. Qed.

forall (x y : R) (P : R -> bool), x <= y -> Un_growing (dicho_lb x y P)

forall (x y : R) (P : R -> bool), x <= y -> Un_growing (dicho_lb x y P)
x, y:R
P:R -> bool
H:x <= y

Un_growing (dicho_lb x y P)
x, y:R
P:R -> bool
H:x <= y

forall n : nat, dicho_lb x y P n <= dicho_lb x y P (S n)
x, y:R
P:R -> bool
H:x <= y
n:nat

dicho_lb x y P n <= dicho_lb x y P (S n)
x, y:R
P:R -> bool
H:x <= y
n:nat

dicho_lb x y P n <= (if P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) then Dichotomy_lb x y P n else (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2)
x, y:R
P:R -> bool
H:x <= y
n:nat

dicho_lb x y P n <= Dichotomy_lb x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
dicho_lb x y P n <= (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2
x, y:R
P:R -> bool
H:x <= y
n:nat

dicho_lb x y P n <= (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2
x, y:R
P:R -> bool
H:x <= y
n:nat

0 < 2
x, y:R
P:R -> bool
H:x <= y
n:nat
2 * dicho_lb x y P n <= 2 * ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2)
x, y:R
P:R -> bool
H:x <= y
n:nat

2 * dicho_lb x y P n <= 2 * ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2)
x, y:R
P:R -> bool
H:x <= y
n:nat

2 * dicho_lb x y P n <= (Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2 * 2
x, y:R
P:R -> bool
H:x <= y
n:nat

2 * dicho_lb x y P n <= (Dichotomy_lb x y P n + Dichotomy_ub x y P n) * 1
x, y:R
P:R -> bool
H:x <= y
n:nat

2 * dicho_lb x y P n <= Dichotomy_lb x y P n + Dichotomy_ub x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat

dicho_lb x y P n + dicho_lb x y P n <= Dichotomy_lb x y P n + Dichotomy_ub x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat

dicho_lb x y P n <= Dichotomy_ub x y P n
replace (Dichotomy_ub x y P n) with (dicho_up x y P n); [ apply dicho_comp; assumption | reflexivity ]. Qed.

forall (x y : R) (P : R -> bool), x <= y -> Un_decreasing (dicho_up x y P)

forall (x y : R) (P : R -> bool), x <= y -> Un_decreasing (dicho_up x y P)
x, y:R
P:R -> bool
H:x <= y

Un_decreasing (dicho_up x y P)
x, y:R
P:R -> bool
H:x <= y

forall n : nat, dicho_up x y P (S n) <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat

dicho_up x y P (S n) <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat

(if P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) then (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 else Dichotomy_ub x y P n) <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat

(Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Dichotomy_ub x y P n <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat

0 < 2
x, y:R
P:R -> bool
H:x <= y
n:nat
2 * ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2) <= 2 * dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Dichotomy_ub x y P n <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat

2 * ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2) <= 2 * dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Dichotomy_ub x y P n <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat

(Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2 * 2 <= 2 * dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Dichotomy_ub x y P n <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat

(Dichotomy_lb x y P n + Dichotomy_ub x y P n) * 1 <= 2 * dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Dichotomy_ub x y P n <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat

Dichotomy_lb x y P n + Dichotomy_ub x y P n <= 2 * dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Dichotomy_ub x y P n <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat

Dichotomy_lb x y P n + Dichotomy_ub x y P n <= dicho_up x y P n + dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Dichotomy_ub x y P n <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat

Dichotomy_lb x y P n + dicho_up x y P n <= dicho_up x y P n + dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Dichotomy_ub x y P n <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat

dicho_lb x y P n + dicho_up x y P n <= dicho_up x y P n + dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Dichotomy_ub x y P n <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat

dicho_up x y P n + dicho_lb x y P n <= dicho_up x y P n + dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Dichotomy_ub x y P n <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat

dicho_lb x y P n <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Dichotomy_ub x y P n <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat

Dichotomy_ub x y P n <= dicho_up x y P n
right; reflexivity. Qed.

forall (x y : R) (P : R -> bool), x <= y -> forall n : nat, dicho_lb x y P n <= y

forall (x y : R) (P : R -> bool), x <= y -> forall n : nat, dicho_lb x y P n <= y
x, y:R
P:R -> bool
H:x <= y
n:nat

dicho_lb x y P n <= y
x, y:R
P:R -> bool
H:x <= y

dicho_lb x y P 0 <= y
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:dicho_lb x y P n <= y
dicho_lb x y P (S n) <= y
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:dicho_lb x y P n <= y

dicho_lb x y P (S n) <= y
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:dicho_lb x y P n <= y

(if P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) then Dichotomy_lb x y P n else (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) <= y
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:dicho_lb x y P n <= y

Dichotomy_lb x y P n <= y
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:dicho_lb x y P n <= y
(Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 <= y
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:dicho_lb x y P n <= y

(Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 <= y
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:dicho_lb x y P n <= y

0 < 2
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:dicho_lb x y P n <= y
2 * ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2) <= 2 * y
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:dicho_lb x y P n <= y

2 * ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2) <= 2 * y
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:dicho_lb x y P n <= y

(Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2 * 2 <= 2 * y
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:dicho_lb x y P n <= y

Dichotomy_lb x y P n + Dichotomy_ub x y P n <= 2 * y
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:dicho_lb x y P n <= y

Dichotomy_lb x y P n <= y
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:dicho_lb x y P n <= y
Dichotomy_ub x y P n <= y
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:dicho_lb x y P n <= y

Dichotomy_ub x y P n <= y
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:dicho_lb x y P n <= y

Dichotomy_ub x y P n <= Dichotomy_ub x y P 0
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:dicho_lb x y P n <= y

Un_decreasing (Dichotomy_ub x y P)
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:dicho_lb x y P n <= y
(0 <= n)%nat
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:dicho_lb x y P n <= y
H0:Un_decreasing (dicho_up x y P)

Un_decreasing (Dichotomy_ub x y P)
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:dicho_lb x y P n <= y
(0 <= n)%nat
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:dicho_lb x y P n <= y

(0 <= n)%nat
apply le_O_n. Qed.

forall (x y : R) (P : R -> bool), x <= y -> has_ub (dicho_lb x y P)

forall (x y : R) (P : R -> bool), x <= y -> has_ub (dicho_lb x y P)
x, y:R
P:R -> bool
H:x <= y

has_ub (dicho_lb x y P)
x, y:R
P:R -> bool
H:x <= y

(forall n : nat, dicho_lb x y P n <= y) -> has_ub (dicho_lb x y P)
x, y:R
P:R -> bool
H:x <= y
forall n : nat, dicho_lb x y P n <= y
x, y:R
P:R -> bool
H:x <= y
H0:forall n : nat, dicho_lb x y P n <= y

has_ub (dicho_lb x y P)
x, y:R
P:R -> bool
H:x <= y
forall n : nat, dicho_lb x y P n <= y
x, y:R
P:R -> bool
H:x <= y
H0:forall n : nat, dicho_lb x y P n <= y

bound (EUn (dicho_lb x y P))
x, y:R
P:R -> bool
H:x <= y
forall n : nat, dicho_lb x y P n <= y
x, y:R
P:R -> bool
H:x <= y
H0:forall n : nat, dicho_lb x y P n <= y

exists m : R, is_upper_bound (EUn (dicho_lb x y P)) m
x, y:R
P:R -> bool
H:x <= y
forall n : nat, dicho_lb x y P n <= y
x, y:R
P:R -> bool
H:x <= y
H0:forall n : nat, dicho_lb x y P n <= y

is_upper_bound (EUn (dicho_lb x y P)) y
x, y:R
P:R -> bool
H:x <= y
forall n : nat, dicho_lb x y P n <= y
x, y:R
P:R -> bool
H:x <= y
H0:forall n : nat, dicho_lb x y P n <= y

forall x0 : R, EUn (dicho_lb x y P) x0 -> x0 <= y
x, y:R
P:R -> bool
H:x <= y
forall n : nat, dicho_lb x y P n <= y
x, y:R
P:R -> bool
H:x <= y
H0:forall n : nat, dicho_lb x y P n <= y
x0:R
H1:EUn (dicho_lb x y P) x0

x0 <= y
x, y:R
P:R -> bool
H:x <= y
forall n : nat, dicho_lb x y P n <= y
x, y:R
P:R -> bool
H:x <= y
H0:forall n : nat, dicho_lb x y P n <= y
x0:R
H1:EUn (dicho_lb x y P) x0
x1:nat
H2:x0 = dicho_lb x y P x1

x0 <= y
x, y:R
P:R -> bool
H:x <= y
forall n : nat, dicho_lb x y P n <= y
x, y:R
P:R -> bool
H:x <= y

forall n : nat, dicho_lb x y P n <= y
apply dicho_lb_maj_y; assumption. Qed.

forall (x y : R) (P : R -> bool), x <= y -> forall n : nat, x <= dicho_up x y P n

forall (x y : R) (P : R -> bool), x <= y -> forall n : nat, x <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat

x <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y

x <= dicho_up x y P 0
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n
x <= dicho_up x y P (S n)
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n

x <= dicho_up x y P (S n)
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n

x <= (if P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) then (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 else Dichotomy_ub x y P n)
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n

x <= (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n
x <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n

0 < 2
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n
2 * x <= 2 * ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2)
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n
x <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n

2 * x <= 2 * ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2)
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n
x <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n

2 * x <= (Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2 * 2
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n
x <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n

2 * x <= Dichotomy_lb x y P n + Dichotomy_ub x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n
x <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n

x <= Dichotomy_lb x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n
x <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n
x <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n

Dichotomy_lb x y P 0 <= Dichotomy_lb x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n
x <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n
x <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n

Un_growing (Dichotomy_lb x y P)
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n
(0 <= n)%nat
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n
x <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n
x <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n
H0:Un_growing (dicho_lb x y P)

Un_growing (Dichotomy_lb x y P)
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n
(0 <= n)%nat
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n
x <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n
x <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n

(0 <= n)%nat
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n
x <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n
x <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n

x <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n
x <= Dichotomy_ub x y P n
x, y:R
P:R -> bool
H:x <= y
n:nat
Hrecn:x <= dicho_up x y P n

x <= Dichotomy_ub x y P n
assumption. Qed.

forall (x y : R) (P : R -> bool), x <= y -> has_lb (dicho_up x y P)

forall (x y : R) (P : R -> bool), x <= y -> has_lb (dicho_up x y P)
x, y:R
P:R -> bool
H:x <= y

has_lb (dicho_up x y P)
x, y:R
P:R -> bool
H:x <= y

(forall n : nat, x <= dicho_up x y P n) -> has_lb (dicho_up x y P)
x, y:R
P:R -> bool
H:x <= y
forall n : nat, x <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
H0:forall n : nat, x <= dicho_up x y P n

has_lb (dicho_up x y P)
x, y:R
P:R -> bool
H:x <= y
forall n : nat, x <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
H0:forall n : nat, x <= dicho_up x y P n

bound (EUn (opp_seq (dicho_up x y P)))
x, y:R
P:R -> bool
H:x <= y
forall n : nat, x <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
H0:forall n : nat, x <= dicho_up x y P n

exists m : R, is_upper_bound (EUn (opp_seq (dicho_up x y P))) m
x, y:R
P:R -> bool
H:x <= y
forall n : nat, x <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
H0:forall n : nat, x <= dicho_up x y P n

is_upper_bound (EUn (opp_seq (dicho_up x y P))) (- x)
x, y:R
P:R -> bool
H:x <= y
forall n : nat, x <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
H0:forall n : nat, x <= dicho_up x y P n

forall x0 : R, EUn (opp_seq (dicho_up x y P)) x0 -> x0 <= - x
x, y:R
P:R -> bool
H:x <= y
forall n : nat, x <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
H0:forall n : nat, x <= dicho_up x y P n
x0:R
H1:EUn (opp_seq (dicho_up x y P)) x0

x0 <= - x
x, y:R
P:R -> bool
H:x <= y
forall n : nat, x <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
H0:forall n : nat, x <= dicho_up x y P n
x0:R
H1:EUn (opp_seq (dicho_up x y P)) x0
x1:nat
H2:x0 = opp_seq (dicho_up x y P) x1

x0 <= - x
x, y:R
P:R -> bool
H:x <= y
forall n : nat, x <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
H0:forall n : nat, x <= dicho_up x y P n
x0:R
H1:EUn (opp_seq (dicho_up x y P)) x0
x1:nat
H2:x0 = opp_seq (dicho_up x y P) x1

opp_seq (dicho_up x y P) x1 <= - x
x, y:R
P:R -> bool
H:x <= y
forall n : nat, x <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
H0:forall n : nat, x <= dicho_up x y P n
x0:R
H1:EUn (opp_seq (dicho_up x y P)) x0
x1:nat
H2:x0 = opp_seq (dicho_up x y P) x1

- dicho_up x y P x1 <= - x
x, y:R
P:R -> bool
H:x <= y
forall n : nat, x <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y
H0:forall n : nat, x <= dicho_up x y P n
x0:R
H1:EUn (opp_seq (dicho_up x y P)) x0
x1:nat
H2:x0 = opp_seq (dicho_up x y P) x1

x <= dicho_up x y P x1
x, y:R
P:R -> bool
H:x <= y
forall n : nat, x <= dicho_up x y P n
x, y:R
P:R -> bool
H:x <= y

forall n : nat, x <= dicho_up x y P n
apply dicho_up_min_x; assumption. Qed.

forall (x y : R) (P : R -> bool), x <= y -> {l : R | Un_cv (dicho_lb x y P) l}

forall (x y : R) (P : R -> bool), x <= y -> {l : R | Un_cv (dicho_lb x y P) l}
x, y:R
P:R -> bool
H:x <= y

{l : R | Un_cv (dicho_lb x y P) l}
x, y:R
P:R -> bool
H:x <= y

Un_growing (dicho_lb x y P)
x, y:R
P:R -> bool
H:x <= y
has_ub (dicho_lb x y P)
x, y:R
P:R -> bool
H:x <= y

has_ub (dicho_lb x y P)
apply dicho_lb_maj; assumption. Qed.

forall (x y : R) (P : R -> bool), x <= y -> {l : R | Un_cv (dicho_up x y P) l}

forall (x y : R) (P : R -> bool), x <= y -> {l : R | Un_cv (dicho_up x y P) l}
x, y:R
P:R -> bool
H:x <= y

{l : R | Un_cv (dicho_up x y P) l}
x, y:R
P:R -> bool
H:x <= y

Un_decreasing (dicho_up x y P)
x, y:R
P:R -> bool
H:x <= y
has_lb (dicho_up x y P)
x, y:R
P:R -> bool
H:x <= y

has_lb (dicho_up x y P)
apply dicho_up_min; assumption. Qed.

forall (x y : R) (P : R -> bool) (n : nat), x <= y -> dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n

forall (x y : R) (P : R -> bool) (n : nat), x <= y -> dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
x, y:R
P:R -> bool
n:nat
H:x <= y

dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
x, y:R
P:R -> bool
H:x <= y

dicho_up x y P 0 - dicho_lb x y P 0 = (y - x) / 2 ^ 0
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
dicho_up x y P (S n) - dicho_lb x y P (S n) = (y - x) / 2 ^ S n
x, y:R
P:R -> bool
H:x <= y

y - x = (y - x) / 1
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
dicho_up x y P (S n) - dicho_lb x y P (S n) = (y - x) / 2 ^ S n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n

dicho_up x y P (S n) - dicho_lb x y P (S n) = (y - x) / 2 ^ S n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n

(if P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) then (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 else Dichotomy_ub x y P n) - (if P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) then Dichotomy_lb x y P n else (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) = (y - x) / (2 * 2 ^ n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n

(Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 - Dichotomy_lb x y P n = (y - x) / (2 * 2 ^ n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
Dichotomy_ub x y P n - (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 = (y - x) / (2 * 2 ^ n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n

(Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2 - Dichotomy_lb x y P n = (y - x) * / (2 * 2 ^ n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
Dichotomy_ub x y P n - (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 = (y - x) / (2 * 2 ^ n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n

(dicho_up x y P n - dicho_lb x y P n) / 2 = (y - x) * / (2 * 2 ^ n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
(dicho_up x y P n - dicho_lb x y P n) / 2 = (Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2 - Dichotomy_lb x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
Dichotomy_ub x y P n - (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 = (y - x) / (2 * 2 ^ n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n

(y - x) / 2 ^ n * / 2 = (y - x) * / (2 * 2 ^ n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
(dicho_up x y P n - dicho_lb x y P n) / 2 = (Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2 - Dichotomy_lb x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
Dichotomy_ub x y P n - (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 = (y - x) / (2 * 2 ^ n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n

(y - x) * / 2 ^ n * / 2 = (y - x) * / (2 * 2 ^ n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
(dicho_up x y P n - dicho_lb x y P n) / 2 = (Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2 - Dichotomy_lb x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
Dichotomy_ub x y P n - (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 = (y - x) / (2 * 2 ^ n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n

(y - x) * / 2 ^ n * / 2 = (y - x) * (/ 2 * / 2 ^ n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
2 <> 0
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
2 ^ n <> 0
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
(dicho_up x y P n - dicho_lb x y P n) / 2 = (Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2 - Dichotomy_lb x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
Dichotomy_ub x y P n - (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 = (y - x) / (2 * 2 ^ n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n

2 <> 0
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
2 ^ n <> 0
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
(dicho_up x y P n - dicho_lb x y P n) / 2 = (Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2 - Dichotomy_lb x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
Dichotomy_ub x y P n - (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 = (y - x) / (2 * 2 ^ n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n

2 ^ n <> 0
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
(dicho_up x y P n - dicho_lb x y P n) / 2 = (Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2 - Dichotomy_lb x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
Dichotomy_ub x y P n - (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 = (y - x) / (2 * 2 ^ n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n

(dicho_up x y P n - dicho_lb x y P n) / 2 = (Dichotomy_lb x y P n + Dichotomy_ub x y P n) * / 2 - Dichotomy_lb x y P n
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
Dichotomy_ub x y P n - (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 = (y - x) / (2 * 2 ^ n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n

Dichotomy_ub x y P n - (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 = (y - x) / (2 * 2 ^ n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n

(dicho_up x y P n - dicho_lb x y P n) / 2 = (y - x) / (2 * 2 ^ n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
(dicho_up x y P n - dicho_lb x y P n) / 2 = Dichotomy_ub x y P n - (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n

(y - x) / 2 ^ n * / 2 = (y - x) * / (2 * 2 ^ n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
(dicho_up x y P n - dicho_lb x y P n) / 2 = Dichotomy_ub x y P n - (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n

(y - x) * / 2 ^ n * / 2 = (y - x) * / (2 * 2 ^ n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
(dicho_up x y P n - dicho_lb x y P n) / 2 = Dichotomy_ub x y P n - (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n

(y - x) * / 2 ^ n * / 2 = (y - x) * (/ 2 * / 2 ^ n)
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
2 <> 0
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
2 ^ n <> 0
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
(dicho_up x y P n - dicho_lb x y P n) / 2 = Dichotomy_ub x y P n - (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n

2 <> 0
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
2 ^ n <> 0
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
(dicho_up x y P n - dicho_lb x y P n) / 2 = Dichotomy_ub x y P n - (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n

2 ^ n <> 0
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n
(dicho_up x y P n - dicho_lb x y P n) / 2 = Dichotomy_ub x y P n - (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2
x, y:R
P:R -> bool
n:nat
H:x <= y
Hrecn:dicho_up x y P n - dicho_lb x y P n = (y - x) / 2 ^ n

(dicho_up x y P n - dicho_lb x y P n) / 2 = Dichotomy_ub x y P n - (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2
pattern (Dichotomy_ub x y P n) at 1; rewrite (double_var (Dichotomy_ub x y P n)); unfold dicho_up, dicho_lb, Rminus, Rdiv; ring. Qed. Definition pow_2_n (n:nat) := 2 ^ n.

forall n : nat, pow_2_n n <> 0

forall n : nat, pow_2_n n <> 0
n:nat

pow_2_n n <> 0
n:nat

2 ^ n <> 0
n:nat

2 <> 0
discrR. Qed.

Un_growing pow_2_n

Un_growing pow_2_n

forall n : nat, pow_2_n n <= pow_2_n (S n)
n:nat

pow_2_n n <= pow_2_n (S n)
n:nat

2 ^ n <= 2 ^ n * 2 ^ 1
n:nat

2 ^ n * 1 <= 2 ^ n * 2 ^ 1
n:nat

0 <= 2 ^ n
n:nat
1 <= 2 ^ 1
n:nat

1 <= 2 ^ 1
n:nat

1 <= 2 * 1
n:nat

1 <= 2
pattern 1 at 1; rewrite <- Rplus_0_r; apply Rplus_le_compat_l; left; apply Rlt_0_1. Qed.

cv_infty pow_2_n

cv_infty pow_2_n

(forall N : nat, INR N <= 2 ^ N) -> cv_infty pow_2_n

forall N : nat, INR N <= 2 ^ N
H:forall N : nat, INR N <= 2 ^ N

cv_infty pow_2_n

forall N : nat, INR N <= 2 ^ N
H:forall N : nat, INR N <= 2 ^ N

forall M : R, exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N : nat, INR N <= 2 ^ N
M:R

exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N : nat, INR N <= 2 ^ N
M:R
Hlt:0 < M

exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z

exists N0 : nat, forall n : nat, (N0 <= n)%nat -> M < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z

(0 <= N)%Z -> exists N0 : nat, forall n : nat, (N0 <= n)%nat -> M < pow_2_n n
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z
(0 <= N)%Z
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z

exists N0 : nat, forall n : nat, (N0 <= n)%nat -> M < pow_2_n n
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z
(0 <= N)%Z
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N1 : nat, INR N1 <= 2 ^ N1
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z
N0:nat
H1:N = Z.of_nat N0

exists N1 : nat, forall n : nat, (N1 <= n)%nat -> M < pow_2_n n
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z
(0 <= N)%Z
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N1 : nat, INR N1 <= 2 ^ N1
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z
N0:nat
H1:N = Z.of_nat N0

forall n : nat, (N0 <= n)%nat -> M < pow_2_n n
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z
(0 <= N)%Z
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N1 : nat, INR N1 <= 2 ^ N1
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z
N0:nat
H1:N = Z.of_nat N0
n:nat
H2:(N0 <= n)%nat

M < pow_2_n n
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z
(0 <= N)%Z
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N1 : nat, INR N1 <= 2 ^ N1
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z
N0:nat
H1:N = Z.of_nat N0
n:nat
H2:(N0 <= n)%nat

M < INR N0
H:forall N1 : nat, INR N1 <= 2 ^ N1
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z
N0:nat
H1:N = Z.of_nat N0
n:nat
H2:(N0 <= n)%nat
INR N0 <= pow_2_n n
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z
(0 <= N)%Z
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N1 : nat, INR N1 <= 2 ^ N1
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z
N0:nat
H1:N = Z.of_nat N0
n:nat
H2:(N0 <= n)%nat

M < IZR (Z.of_nat N0)
H:forall N1 : nat, INR N1 <= 2 ^ N1
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z
N0:nat
H1:N = Z.of_nat N0
n:nat
H2:(N0 <= n)%nat
INR N0 <= pow_2_n n
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z
(0 <= N)%Z
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N1 : nat, INR N1 <= 2 ^ N1
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z
N0:nat
H1:N = Z.of_nat N0
n:nat
H2:(N0 <= n)%nat

M < IZR N
H:forall N1 : nat, INR N1 <= 2 ^ N1
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z
N0:nat
H1:N = Z.of_nat N0
n:nat
H2:(N0 <= n)%nat
INR N0 <= pow_2_n n
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z
(0 <= N)%Z
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N1 : nat, INR N1 <= 2 ^ N1
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z
N0:nat
H1:N = Z.of_nat N0
n:nat
H2:(N0 <= n)%nat

M < IZR (up M)
H:forall N1 : nat, INR N1 <= 2 ^ N1
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z
N0:nat
H1:N = Z.of_nat N0
n:nat
H2:(N0 <= n)%nat
INR N0 <= pow_2_n n
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z
(0 <= N)%Z
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N1 : nat, INR N1 <= 2 ^ N1
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z
N0:nat
H1:N = Z.of_nat N0
n:nat
H2:(N0 <= n)%nat
H3:IZR (up M) > M /\ IZR (up M) - M <= 1

M < IZR (up M)
H:forall N1 : nat, INR N1 <= 2 ^ N1
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z
N0:nat
H1:N = Z.of_nat N0
n:nat
H2:(N0 <= n)%nat
INR N0 <= pow_2_n n
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z
(0 <= N)%Z
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N1 : nat, INR N1 <= 2 ^ N1
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z
N0:nat
H1:N = Z.of_nat N0
n:nat
H2:(N0 <= n)%nat

INR N0 <= pow_2_n n
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z
(0 <= N)%Z
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N1 : nat, INR N1 <= 2 ^ N1
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z
N0:nat
H1:N = Z.of_nat N0
n:nat
H2:(N0 <= n)%nat

INR N0 <= pow_2_n N0
H:forall N1 : nat, INR N1 <= 2 ^ N1
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z
N0:nat
H1:N = Z.of_nat N0
n:nat
H2:(N0 <= n)%nat
pow_2_n N0 <= pow_2_n n
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z
(0 <= N)%Z
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N1 : nat, INR N1 <= 2 ^ N1
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z
N0:nat
H1:N = Z.of_nat N0
n:nat
H2:(N0 <= n)%nat

pow_2_n N0 <= pow_2_n n
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z
(0 <= N)%Z
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N1 : nat, INR N1 <= 2 ^ N1
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z
N0:nat
H1:N = Z.of_nat N0
n:nat
H2:(N0 <= n)%nat

pow_2_n n >= pow_2_n N0
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z
(0 <= N)%Z
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N1 : nat, INR N1 <= 2 ^ N1
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z
N0:nat
H1:N = Z.of_nat N0
n:nat
H2:(N0 <= n)%nat

Un_growing pow_2_n
H:forall N1 : nat, INR N1 <= 2 ^ N1
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z
N0:nat
H1:N = Z.of_nat N0
n:nat
H2:(N0 <= n)%nat
(n >= N0)%nat
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z
(0 <= N)%Z
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N1 : nat, INR N1 <= 2 ^ N1
M:R
Hlt:0 < M
N:=up M:Z
H0:(0 <= N)%Z
N0:nat
H1:N = Z.of_nat N0
n:nat
H2:(N0 <= n)%nat

(n >= N0)%nat
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z
(0 <= N)%Z
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z

(0 <= N)%Z
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z

0 <= IZR N
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z

0 <= IZR (up M)
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z

0 <= IZR (up M)
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N0 : nat, INR N0 <= 2 ^ N0
M:R
Hlt:0 < M
N:=up M:Z
H0:IZR (up M) > M /\ IZR (up M) - M <= 1
H1:IZR (up M) > M
H2:IZR (up M) - M <= 1

0 <= IZR (up M)
H:forall N : nat, INR N <= 2 ^ N
exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N : nat, INR N <= 2 ^ N

exists N : nat, forall n : nat, (N <= n)%nat -> 0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N : nat, INR N <= 2 ^ N
n:nat
H0:(0 <= n)%nat

0 < pow_2_n n
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M

exists N : nat, forall n : nat, (N <= n)%nat -> M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
n:nat
H0:(0 <= n)%nat

M < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
n:nat
H0:(0 <= n)%nat

M < 0
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
n:nat
H0:(0 <= n)%nat
0 < pow_2_n n

forall N : nat, INR N <= 2 ^ N
H:forall N : nat, INR N <= 2 ^ N
M:R
Hgt:0 > M
n:nat
H0:(0 <= n)%nat

0 < pow_2_n n

forall N : nat, INR N <= 2 ^ N

forall N : nat, INR N <= 2 ^ N
N:nat

INR 0 <= 2 ^ 0
N:nat
forall n : nat, INR n <= 2 ^ n -> INR (S n) <= 2 ^ S n
N:nat

0 <= 1
N:nat
forall n : nat, INR n <= 2 ^ n -> INR (S n) <= 2 ^ S n
N:nat

forall n : nat, INR n <= 2 ^ n -> INR (S n) <= 2 ^ S n
N, n:nat
H:INR n <= 2 ^ n

INR (S n) <= 2 ^ S n
N, n:nat
H:INR n <= 2 ^ n

INR (S n) <= 2 ^ (n + 1)
N, n:nat
H:INR n <= 2 ^ n

INR n + 1 <= 2 ^ n * 2 ^ 1
N, n:nat
H:INR n <= 2 ^ n

INR n + 1 <= 2 ^ n * (2 * 1)
N, n:nat
H:INR n <= 2 ^ n

INR n + 1 <= 2 ^ n * 2
N, n:nat
H:INR n <= 2 ^ n

INR n + 1 <= 2 ^ n
N, n:nat
H:INR n <= 2 ^ n
2 ^ n <= 2 ^ n * 2
N, n:nat
H:INR n <= 2 ^ n

1 + INR n <= 2 ^ n
N, n:nat
H:INR n <= 2 ^ n
2 ^ n <= 2 ^ n * 2
N, n:nat
H:INR n <= 2 ^ n

1 + INR n * 1 <= 2 ^ n
N, n:nat
H:INR n <= 2 ^ n
2 ^ n <= 2 ^ n * 2
N, n:nat
H:INR n <= 2 ^ n

0 < 1
N, n:nat
H:INR n <= 2 ^ n
2 ^ n <= 2 ^ n * 2
N, n:nat
H:INR n <= 2 ^ n

2 ^ n <= 2 ^ n * 2
N, n:nat
H:INR n <= 2 ^ n

2 ^ n + 0 <= 2 ^ n * 2
N, n:nat
H:INR n <= 2 ^ n

2 ^ n + 0 <= 2 * 2 ^ n
N, n:nat
H:INR n <= 2 ^ n

2 ^ n + 0 <= 2 ^ n + 2 ^ n
N, n:nat
H:INR n <= 2 ^ n

0 <= 2 ^ n
left; apply pow_lt; prove_sup0. Qed.

forall (x y l1 l2 : R) (P : R -> bool), x <= y -> Un_cv (dicho_lb x y P) l1 -> Un_cv (dicho_up x y P) l2 -> l1 = l2

forall (x y l1 l2 : R) (P : R -> bool), x <= y -> Un_cv (dicho_lb x y P) l1 -> Un_cv (dicho_up x y P) l2 -> l1 = l2
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2

l1 = l2
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)

l1 = l2
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)

Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) 0 -> l1 = l2
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) 0
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
H3:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) 0

l1 = l2
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) 0
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
H3:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) 0
H4:l1 - l2 = 0

l1 = l2
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) 0
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)

Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) 0
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)

forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0

exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0

exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hlt:x < y

exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y

exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y

0 < y - x -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x

exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x

0 < eps / (y - x) -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
0 < eps / (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)

exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
0 < eps / (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps / (y - x)

exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
0 < eps / (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat

Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
0 < eps / (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat

Rabs (dicho_lb x y P n - dicho_up x y P n) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
0 < eps / (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat

Rabs (- (dicho_lb x y P n - dicho_up x y P n)) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
0 < eps / (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat

Rabs (dicho_up x y P n - dicho_lb x y P n) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
0 < eps / (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat

Rabs ((y - x) / 2 ^ n) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat
x <= y
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
0 < eps / (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat

Rabs (y - x) * Rabs (/ 2 ^ n) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat
x <= y
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
0 < eps / (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat

(y - x) * Rabs (/ 2 ^ n) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat
y - x >= 0
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat
x <= y
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
0 < eps / (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat

0 < / (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat
/ (y - x) * ((y - x) * Rabs (/ 2 ^ n)) < / (y - x) * eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat
y - x >= 0
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat
x <= y
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
0 < eps / (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat

/ (y - x) * ((y - x) * Rabs (/ 2 ^ n)) < / (y - x) * eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat
y - x >= 0
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat
x <= y
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
0 < eps / (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat

1 * Rabs (/ 2 ^ n) < / (y - x) * eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat
y - x <> 0
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat
y - x >= 0
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat
x <= y
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
0 < eps / (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat

Rabs (/ 2 ^ n) < / (y - x) * eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat
y - x <> 0
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat
y - x >= 0
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat
x <= y
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
0 < eps / (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat

y - x <> 0
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat
y - x >= 0
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat
x <= y
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
0 < eps / (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat

y - x >= 0
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat
x <= y
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
0 < eps / (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat

0 <= y - x
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat
x <= y
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
0 < eps / (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat

x <= x + (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat
x <= y
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
0 < eps / (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (/ pow_2_n n0 - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
H5:0 < eps / (y - x)
N:nat
H6:forall n0 : nat, (n0 >= N)%nat -> Rabs (/ pow_2_n n0 - 0) < eps / (y - x)
n:nat
H7:(n >= N)%nat

x <= y
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x
0 < eps / (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
Hyp:0 < y - x

0 < eps / (y - x)
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y
0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y

0 < y - x
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (/ pow_2_n n - 0) < eps0
Hlt:x < y

x < x + (y - x)
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0

exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n0 : nat => / pow_2_n n0) 0
n:nat
H5:(n >= 0)%nat

Rabs (dicho_lb y y P n - dicho_up y y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n0 : nat => / pow_2_n n0) 0
n:nat
H5:(n >= 0)%nat

Rabs (dicho_lb y y P n - dicho_up y y P n) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n0 : nat => / pow_2_n n0) 0
n:nat
H5:(n >= 0)%nat

Rabs (- (dicho_lb y y P n - dicho_up y y P n)) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n0 : nat => / pow_2_n n0) 0
n:nat
H5:(n >= 0)%nat

Rabs (dicho_up y y P n - dicho_lb y y P n) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n0 : nat => / pow_2_n n0) 0
n:nat
H5:(n >= 0)%nat

Rabs ((y - y) / 2 ^ n) < eps
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n0 : nat => / pow_2_n n0) 0
n:nat
H5:(n >= 0)%nat
y <= y
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
y, l1, l2:R
P:R -> bool
H2:Un_cv (fun i : nat => dicho_lb y y P i - dicho_up y y P i) (l1 - l2)
H1:Un_cv (dicho_up y y P) l2
H0:Un_cv (dicho_lb y y P) l1
H:y <= y
eps:R
H3:eps > 0
H4:Un_cv (fun n0 : nat => / pow_2_n n0) 0
n:nat
H5:(n >= 0)%nat

y <= y
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y
exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
x, y, l1, l2:R
P:R -> bool
H:x <= y
H0:Un_cv (dicho_lb x y P) l1
H1:Un_cv (dicho_up x y P) l2
H2:Un_cv (fun i : nat => dicho_lb x y P i - dicho_up x y P i) (l1 - l2)
eps:R
H3:eps > 0
H4:Un_cv (fun n : nat => / pow_2_n n) 0
Hgt:x > y

exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (dicho_lb x y P n - dicho_up x y P n - 0) < eps
elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H Hgt)). Qed. Definition cond_positivity (x:R) : bool := match Rle_dec 0 x with | left _ => true | right _ => false end.
Sequential characterisation of continuity

forall (f : R -> R) (Un : nat -> R) (l : R), continuity_pt f l -> Un_cv Un l -> Un_cv (fun i : nat => f (Un i)) (f l)

forall (f : R -> R) (Un : nat -> R) (l : R), continuity_pt f l -> Un_cv Un l -> Un_cv (fun i : nat => f (Un i)) (f l)

forall (f : R -> R) (Un : nat -> R) (l : R), limit1_in f (D_x no_cond l) (f l) l -> (forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Un n) l < eps) -> forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (f (Un n)) (f l) < eps

forall (f : R -> R) (Un : nat -> R) (l : R), limit_in R_met R_met f (D_x no_cond l) l (f l) -> (forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Un n) l < eps) -> forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (f (Un n)) (f l) < eps

forall (f : R -> R) (Un : nat -> R) (l : R), (forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x : Base R_met, D_x no_cond l x /\ dist R_met x l < alp -> dist R_met (f x) (f l) < eps)) -> (forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Un n) l < eps) -> forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (f (Un n)) (f l) < eps

forall (f : R -> R) (Un : nat -> R) (l : R), (forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x : Base R_met, D_x no_cond l x /\ (let (Base, dist, _, _, _, _) as m return (Base m -> Base m -> R) := R_met in dist) x l < alp -> (let (Base, dist, _, _, _, _) as m return (Base m -> Base m -> R) := R_met in dist) (f x) (f l) < eps)) -> (forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Un n) l < eps) -> forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (f (Un n)) (f l) < eps

forall (f : R -> R) (Un : nat -> R) (l : R), (forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x : R, D_x no_cond l x /\ R_dist x l < alp -> R_dist (f x) (f l) < eps)) -> (forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Un n) l < eps) -> forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (f (Un n)) (f l) < eps

forall (f : R -> R) (Un : nat -> R) (l : R), (forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps)) -> (forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (Un n - l) < eps) -> forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Un n) - f l) < eps
f:R -> R
Un:nat -> R
l:R
H:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps0)
H0:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (Un n - l) < eps0
eps:R
H1:eps > 0

exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Un n) - f l) < eps
f:R -> R
Un:nat -> R
l:R
H:forall eps0 : R, eps0 > 0 -> exists alp0 : R, alp0 > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp0 -> Rabs (f x - f l) < eps0)
H0:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (Un n - l) < eps0
eps:R
H1:eps > 0
alp:R
H2:alp > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps)

exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Un n) - f l) < eps
f:R -> R
Un:nat -> R
l:R
H:forall eps0 : R, eps0 > 0 -> exists alp0 : R, alp0 > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp0 -> Rabs (f x - f l) < eps0)
H0:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (Un n - l) < eps0
eps:R
H1:eps > 0
alp:R
H2:alp > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps)
H3:alp > 0
H4:forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps

exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Un n) - f l) < eps
f:R -> R
Un:nat -> R
l:R
H:forall eps0 : R, eps0 > 0 -> exists alp0 : R, alp0 > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp0 -> Rabs (f x - f l) < eps0)
H0:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (Un n - l) < eps0
eps:R
H1:eps > 0
alp:R
H2:alp > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps)
H3:alp > 0
H4:forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps
N:nat
H5:forall n : nat, (n >= N)%nat -> Rabs (Un n - l) < alp

exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (f (Un n) - f l) < eps
f:R -> R
Un:nat -> R
l:R
H:forall eps0 : R, eps0 > 0 -> exists alp0 : R, alp0 > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp0 -> Rabs (f x - f l) < eps0)
H0:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - l) < eps0
eps:R
H1:eps > 0
alp:R
H2:alp > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps)
H3:alp > 0
H4:forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps
N:nat
H5:forall n0 : nat, (n0 >= N)%nat -> Rabs (Un n0 - l) < alp
n:nat
H6:(n >= N)%nat

Rabs (f (Un n) - f l) < eps
f:R -> R
Un:nat -> R
l:R
H:forall eps0 : R, eps0 > 0 -> exists alp0 : R, alp0 > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp0 -> Rabs (f x - f l) < eps0)
H0:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - l) < eps0
eps:R
H1:eps > 0
alp:R
H2:alp > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps)
H3:alp > 0
H4:forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps
N:nat
H5:forall n0 : nat, (n0 >= N)%nat -> Rabs (Un n0 - l) < alp
n:nat
H6:(n >= N)%nat
H7:Un n = l

Rabs (f (Un n) - f l) < eps
f:R -> R
Un:nat -> R
l:R
H:forall eps0 : R, eps0 > 0 -> exists alp0 : R, alp0 > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp0 -> Rabs (f x - f l) < eps0)
H0:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - l) < eps0
eps:R
H1:eps > 0
alp:R
H2:alp > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps)
H3:alp > 0
H4:forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps
N:nat
H5:forall n0 : nat, (n0 >= N)%nat -> Rabs (Un n0 - l) < alp
n:nat
H6:(n >= N)%nat
H7:Un n <> l
Rabs (f (Un n) - f l) < eps
f:R -> R
Un:nat -> R
l:R
H:forall eps0 : R, eps0 > 0 -> exists alp0 : R, alp0 > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp0 -> Rabs (f x - f l) < eps0)
H0:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - l) < eps0
eps:R
H1:eps > 0
alp:R
H2:alp > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps)
H3:alp > 0
H4:forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps
N:nat
H5:forall n0 : nat, (n0 >= N)%nat -> Rabs (Un n0 - l) < alp
n:nat
H6:(n >= N)%nat
H7:Un n <> l

Rabs (f (Un n) - f l) < eps
f:R -> R
Un:nat -> R
l:R
H:forall eps0 : R, eps0 > 0 -> exists alp0 : R, alp0 > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp0 -> Rabs (f x - f l) < eps0)
H0:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - l) < eps0
eps:R
H1:eps > 0
alp:R
H2:alp > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps)
H3:alp > 0
H4:forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps
N:nat
H5:forall n0 : nat, (n0 >= N)%nat -> Rabs (Un n0 - l) < alp
n:nat
H6:(n >= N)%nat
H7:Un n <> l

D_x no_cond l (Un n) /\ Rabs (Un n - l) < alp
f:R -> R
Un:nat -> R
l:R
H:forall eps0 : R, eps0 > 0 -> exists alp0 : R, alp0 > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp0 -> Rabs (f x - f l) < eps0)
H0:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - l) < eps0
eps:R
H1:eps > 0
alp:R
H2:alp > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps)
H3:alp > 0
H4:forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps
N:nat
H5:forall n0 : nat, (n0 >= N)%nat -> Rabs (Un n0 - l) < alp
n:nat
H6:(n >= N)%nat
H7:Un n <> l

D_x no_cond l (Un n)
f:R -> R
Un:nat -> R
l:R
H:forall eps0 : R, eps0 > 0 -> exists alp0 : R, alp0 > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp0 -> Rabs (f x - f l) < eps0)
H0:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - l) < eps0
eps:R
H1:eps > 0
alp:R
H2:alp > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps)
H3:alp > 0
H4:forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps
N:nat
H5:forall n0 : nat, (n0 >= N)%nat -> Rabs (Un n0 - l) < alp
n:nat
H6:(n >= N)%nat
H7:Un n <> l
Rabs (Un n - l) < alp
f:R -> R
Un:nat -> R
l:R
H:forall eps0 : R, eps0 > 0 -> exists alp0 : R, alp0 > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp0 -> Rabs (f x - f l) < eps0)
H0:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - l) < eps0
eps:R
H1:eps > 0
alp:R
H2:alp > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps)
H3:alp > 0
H4:forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps
N:nat
H5:forall n0 : nat, (n0 >= N)%nat -> Rabs (Un n0 - l) < alp
n:nat
H6:(n >= N)%nat
H7:Un n <> l

True /\ l <> Un n
f:R -> R
Un:nat -> R
l:R
H:forall eps0 : R, eps0 > 0 -> exists alp0 : R, alp0 > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp0 -> Rabs (f x - f l) < eps0)
H0:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - l) < eps0
eps:R
H1:eps > 0
alp:R
H2:alp > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps)
H3:alp > 0
H4:forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps
N:nat
H5:forall n0 : nat, (n0 >= N)%nat -> Rabs (Un n0 - l) < alp
n:nat
H6:(n >= N)%nat
H7:Un n <> l
Rabs (Un n - l) < alp
f:R -> R
Un:nat -> R
l:R
H:forall eps0 : R, eps0 > 0 -> exists alp0 : R, alp0 > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp0 -> Rabs (f x - f l) < eps0)
H0:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - l) < eps0
eps:R
H1:eps > 0
alp:R
H2:alp > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps)
H3:alp > 0
H4:forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps
N:nat
H5:forall n0 : nat, (n0 >= N)%nat -> Rabs (Un n0 - l) < alp
n:nat
H6:(n >= N)%nat
H7:Un n <> l

True
f:R -> R
Un:nat -> R
l:R
H:forall eps0 : R, eps0 > 0 -> exists alp0 : R, alp0 > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp0 -> Rabs (f x - f l) < eps0)
H0:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - l) < eps0
eps:R
H1:eps > 0
alp:R
H2:alp > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps)
H3:alp > 0
H4:forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps
N:nat
H5:forall n0 : nat, (n0 >= N)%nat -> Rabs (Un n0 - l) < alp
n:nat
H6:(n >= N)%nat
H7:Un n <> l
l <> Un n
f:R -> R
Un:nat -> R
l:R
H:forall eps0 : R, eps0 > 0 -> exists alp0 : R, alp0 > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp0 -> Rabs (f x - f l) < eps0)
H0:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - l) < eps0
eps:R
H1:eps > 0
alp:R
H2:alp > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps)
H3:alp > 0
H4:forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps
N:nat
H5:forall n0 : nat, (n0 >= N)%nat -> Rabs (Un n0 - l) < alp
n:nat
H6:(n >= N)%nat
H7:Un n <> l
Rabs (Un n - l) < alp
f:R -> R
Un:nat -> R
l:R
H:forall eps0 : R, eps0 > 0 -> exists alp0 : R, alp0 > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp0 -> Rabs (f x - f l) < eps0)
H0:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - l) < eps0
eps:R
H1:eps > 0
alp:R
H2:alp > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps)
H3:alp > 0
H4:forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps
N:nat
H5:forall n0 : nat, (n0 >= N)%nat -> Rabs (Un n0 - l) < alp
n:nat
H6:(n >= N)%nat
H7:Un n <> l

l <> Un n
f:R -> R
Un:nat -> R
l:R
H:forall eps0 : R, eps0 > 0 -> exists alp0 : R, alp0 > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp0 -> Rabs (f x - f l) < eps0)
H0:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - l) < eps0
eps:R
H1:eps > 0
alp:R
H2:alp > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps)
H3:alp > 0
H4:forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps
N:nat
H5:forall n0 : nat, (n0 >= N)%nat -> Rabs (Un n0 - l) < alp
n:nat
H6:(n >= N)%nat
H7:Un n <> l
Rabs (Un n - l) < alp
f:R -> R
Un:nat -> R
l:R
H:forall eps0 : R, eps0 > 0 -> exists alp0 : R, alp0 > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp0 -> Rabs (f x - f l) < eps0)
H0:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - l) < eps0
eps:R
H1:eps > 0
alp:R
H2:alp > 0 /\ (forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps)
H3:alp > 0
H4:forall x : R, D_x no_cond l x /\ Rabs (x - l) < alp -> Rabs (f x - f l) < eps
N:nat
H5:forall n0 : nat, (n0 >= N)%nat -> Rabs (Un n0 - l) < alp
n:nat
H6:(n >= N)%nat
H7:Un n <> l

Rabs (Un n - l) < alp
apply H5; assumption. Qed.

forall (x y : R) (P : R -> bool) (n : nat), P x = false -> P (dicho_lb x y P n) = false

forall (x y : R) (P : R -> bool) (n : nat), P x = false -> P (dicho_lb x y P n) = false
x, y:R
P:R -> bool
n:nat
H:P x = false

P (dicho_lb x y P n) = false
x, y:R
P:R -> bool
H:P x = false

P (dicho_lb x y P 0) = false
x, y:R
P:R -> bool
n:nat
H:P x = false
Hrecn:P (dicho_lb x y P n) = false
P (dicho_lb x y P (S n)) = false
x, y:R
P:R -> bool
H:P x = false

P (dicho_lb x y P 0) = false
assumption.
x, y:R
P:R -> bool
n:nat
H:P x = false
Hrecn:P (dicho_lb x y P n) = false

P (dicho_lb x y P (S n)) = false
x, y:R
P:R -> bool
n:nat
H:P x = false
Hrecn:P (dicho_lb x y P n) = false

P (if P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) then Dichotomy_lb x y P n else (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) = false
x, y:R
P:R -> bool
n:nat
H:P x = false
Hrecn:P (dicho_lb x y P n) = false
Heq:P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) = true

P (if P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) then Dichotomy_lb x y P n else (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) = false
x, y:R
P:R -> bool
n:nat
H:P x = false
Hrecn:P (dicho_lb x y P n) = false
Heq:P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) = false
P (if P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) then Dichotomy_lb x y P n else (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) = false
x, y:R
P:R -> bool
n:nat
H:P x = false
Hrecn:P (dicho_lb x y P n) = false
Heq:P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) = true

P (if P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) then Dichotomy_lb x y P n else (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) = false
x, y:R
P:R -> bool
n:nat
H:P x = false
Hrecn:P (dicho_lb x y P n) = false
Heq:P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) = true

P (Dichotomy_lb x y P n) = false
unfold dicho_lb in Hrecn; assumption.
x, y:R
P:R -> bool
n:nat
H:P x = false
Hrecn:P (dicho_lb x y P n) = false
Heq:P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) = false

P (if P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) then Dichotomy_lb x y P n else (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) = false
x, y:R
P:R -> bool
n:nat
H:P x = false
Hrecn:P (dicho_lb x y P n) = false
Heq:P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) = false

P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) = false
assumption. Qed.

forall (x y : R) (P : R -> bool) (n : nat), P y = true -> P (dicho_up x y P n) = true

forall (x y : R) (P : R -> bool) (n : nat), P y = true -> P (dicho_up x y P n) = true
x, y:R
P:R -> bool
n:nat
H:P y = true

P (dicho_up x y P n) = true
x, y:R
P:R -> bool
H:P y = true

P (dicho_up x y P 0) = true
x, y:R
P:R -> bool
n:nat
H:P y = true
Hrecn:P (dicho_up x y P n) = true
P (dicho_up x y P (S n)) = true
x, y:R
P:R -> bool
H:P y = true

P (dicho_up x y P 0) = true
assumption.
x, y:R
P:R -> bool
n:nat
H:P y = true
Hrecn:P (dicho_up x y P n) = true

P (dicho_up x y P (S n)) = true
x, y:R
P:R -> bool
n:nat
H:P y = true
Hrecn:P (dicho_up x y P n) = true

P (if P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) then (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 else Dichotomy_ub x y P n) = true
x, y:R
P:R -> bool
n:nat
H:P y = true
Hrecn:P (dicho_up x y P n) = true
Heq:P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) = true

P (if P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) then (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 else Dichotomy_ub x y P n) = true
x, y:R
P:R -> bool
n:nat
H:P y = true
Hrecn:P (dicho_up x y P n) = true
Heq:P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) = false
P (if P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) then (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 else Dichotomy_ub x y P n) = true
x, y:R
P:R -> bool
n:nat
H:P y = true
Hrecn:P (dicho_up x y P n) = true
Heq:P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) = true

P (if P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) then (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 else Dichotomy_ub x y P n) = true
x, y:R
P:R -> bool
n:nat
H:P y = true
Hrecn:P (dicho_up x y P n) = true
Heq:P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) = true

P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) = true
unfold dicho_lb in Hrecn; assumption.
x, y:R
P:R -> bool
n:nat
H:P y = true
Hrecn:P (dicho_up x y P n) = true
Heq:P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) = false

P (if P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) then (Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2 else Dichotomy_ub x y P n) = true
x, y:R
P:R -> bool
n:nat
H:P y = true
Hrecn:P (dicho_up x y P n) = true
Heq:P ((Dichotomy_lb x y P n + Dichotomy_ub x y P n) / 2) = false

P (Dichotomy_ub x y P n) = true
assumption. Qed. (* A general purpose corollary. *)

forall a : R, Un_cv (fun n : nat => a / 2 ^ n) 0
a:R

Un_cv (fun n : nat => a * / 2 ^ n) (a * 0)
a:R

Un_cv (fun _ : nat => a) a
a:R
Un_cv (fun i : nat => / 2 ^ i) 0
a:R

Un_cv (fun i : nat => / 2 ^ i) 0
exact (cv_infty_cv_R0 pow_2_n pow_2_n_neq_R0 pow_2_n_infty). Qed.
Intermediate Value Theorem

forall (f : R -> R) (x y : R), continuity f -> x < y -> f x < 0 -> 0 < f y -> {z : R | x <= z <= y /\ f z = 0}

forall (f : R -> R) (x y : R), continuity f -> x < y -> f x < 0 -> 0 < f y -> {z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x1

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x1
x0:R
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x1
x0:R
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0

x <= x0 <= y /\ f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0

x <= x0 <= y
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0

x <= x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
x0 <= y
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0

x <= dicho_lb x y (fun z : R => cond_positivity (f z)) 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
dicho_lb x y (fun z : R => cond_positivity (f z)) 0 <= x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
x0 <= y
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0

x <= x
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
dicho_lb x y (fun z : R => cond_positivity (f z)) 0 <= x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
x0 <= y
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0

dicho_lb x y (fun z : R => cond_positivity (f z)) 0 <= x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
x0 <= y
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0

Un_growing (dicho_lb x y (fun z : R => cond_positivity (f z)))
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
x0 <= y
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0

Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
x0 <= y
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0

x0 <= y
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0

x0 <= dicho_up x y (fun z : R => cond_positivity (f z)) 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
dicho_up x y (fun z : R => cond_positivity (f z)) 0 <= y
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0

Un_decreasing (dicho_up x y (fun z : R => cond_positivity (f z)))
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
dicho_up x y (fun z : R => cond_positivity (f z)) 0 <= y
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0

Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
dicho_up x y (fun z : R => cond_positivity (f z)) 0 <= y
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0

dicho_up x y (fun z : R => cond_positivity (f z)) 0 <= y
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0

f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R

f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R

f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R

((forall n : nat, f (Vn n) <= 0) -> f x0 <= 0) -> f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R

((forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0) -> ((forall n : nat, f (Vn n) <= 0) -> f x0 <= 0) -> f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0

f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0

(forall n : nat, f (Vn n) <= 0) -> f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0

(forall n : nat, 0 <= f (Wn n)) -> (forall n : nat, f (Vn n) <= 0) -> f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, 0 <= f (Wn n)
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
H7:forall n : nat, 0 <= f (Wn n)
H8:forall n : nat, f (Vn n) <= 0

f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, 0 <= f (Wn n)
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
H7:forall n : nat, 0 <= f (Wn n)
H8:forall n : nat, f (Vn n) <= 0
H9:f x0 <= 0

f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, 0 <= f (Wn n)
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
H7:forall n : nat, 0 <= f (Wn n)
H8:forall n : nat, f (Vn n) <= 0
H9:f x0 <= 0
H10:0 <= f x0

f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, 0 <= f (Wn n)
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0

forall n : nat, 0 <= f (Wn n)
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat

0 <= f (Wn n)
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat

0 <= f (dicho_up x y (fun z : R => cond_positivity (f z)) n)
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat

(forall z : R, cond_positivity z = true <-> 0 <= z) -> 0 <= f (dicho_up x y (fun z : R => cond_positivity (f z)) n)
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
forall z : R, cond_positivity z = true <-> 0 <= z
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
H7:forall z : R, cond_positivity z = true <-> 0 <= z

0 <= f (dicho_up x y (fun z : R => cond_positivity (f z)) n)
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
forall z : R, cond_positivity z = true <-> 0 <= z
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
H7:forall z : R, cond_positivity z = true <-> 0 <= z
H8:(fun z : R => cond_positivity (f z)) y = true -> (fun z : R => cond_positivity (f z)) (dicho_up x y (fun z : R => cond_positivity (f z)) n) = true

0 <= f (dicho_up x y (fun z : R => cond_positivity (f z)) n)
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
forall z : R, cond_positivity z = true <-> 0 <= z
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
H7:forall z : R, cond_positivity z = true <-> 0 <= z
H8:(fun z : R => cond_positivity (f z)) y = true -> (fun z : R => cond_positivity (f z)) (dicho_up x y (fun z : R => cond_positivity (f z)) n) = true
H9:cond_positivity (f (dicho_up x y (fun z : R => cond_positivity (f z)) n)) = true -> 0 <= f (dicho_up x y (fun z : R => cond_positivity (f z)) n)
H10:0 <= f (dicho_up x y (fun z : R => cond_positivity (f z)) n) -> cond_positivity (f (dicho_up x y (fun z : R => cond_positivity (f z)) n)) = true

0 <= f (dicho_up x y (fun z : R => cond_positivity (f z)) n)
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
forall z : R, cond_positivity z = true <-> 0 <= z
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
H7:forall z : R, cond_positivity z = true <-> 0 <= z
H8:(fun z : R => cond_positivity (f z)) y = true -> (fun z : R => cond_positivity (f z)) (dicho_up x y (fun z : R => cond_positivity (f z)) n) = true
H9:cond_positivity (f (dicho_up x y (fun z : R => cond_positivity (f z)) n)) = true -> 0 <= f (dicho_up x y (fun z : R => cond_positivity (f z)) n)
H10:0 <= f (dicho_up x y (fun z : R => cond_positivity (f z)) n) -> cond_positivity (f (dicho_up x y (fun z : R => cond_positivity (f z)) n)) = true

cond_positivity (f (dicho_up x y (fun z : R => cond_positivity (f z)) n)) = true
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
forall z : R, cond_positivity z = true <-> 0 <= z
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
H7:forall z : R, cond_positivity z = true <-> 0 <= z
H8:(fun z : R => cond_positivity (f z)) y = true -> (fun z : R => cond_positivity (f z)) (dicho_up x y (fun z : R => cond_positivity (f z)) n) = true
H9:cond_positivity (f (dicho_up x y (fun z : R => cond_positivity (f z)) n)) = true -> 0 <= f (dicho_up x y (fun z : R => cond_positivity (f z)) n)
H10:0 <= f (dicho_up x y (fun z : R => cond_positivity (f z)) n) -> cond_positivity (f (dicho_up x y (fun z : R => cond_positivity (f z)) n)) = true

cond_positivity (f y) = true
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
forall z : R, cond_positivity z = true <-> 0 <= z
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
H7:forall z : R, cond_positivity z = true <-> 0 <= z
H8:(fun z : R => cond_positivity (f z)) y = true -> (fun z : R => cond_positivity (f z)) (dicho_up x y (fun z : R => cond_positivity (f z)) n) = true
H9:cond_positivity (f (dicho_up x y (fun z : R => cond_positivity (f z)) n)) = true -> 0 <= f (dicho_up x y (fun z : R => cond_positivity (f z)) n)
H10:0 <= f (dicho_up x y (fun z : R => cond_positivity (f z)) n) -> cond_positivity (f (dicho_up x y (fun z : R => cond_positivity (f z)) n)) = true
H11:cond_positivity (f y) = true -> 0 <= f y
H12:0 <= f y -> cond_positivity (f y) = true

cond_positivity (f y) = true
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
forall z : R, cond_positivity z = true <-> 0 <= z
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
H7:forall z : R, cond_positivity z = true <-> 0 <= z
H8:(fun z : R => cond_positivity (f z)) y = true -> (fun z : R => cond_positivity (f z)) (dicho_up x y (fun z : R => cond_positivity (f z)) n) = true
H9:cond_positivity (f (dicho_up x y (fun z : R => cond_positivity (f z)) n)) = true -> 0 <= f (dicho_up x y (fun z : R => cond_positivity (f z)) n)
H10:0 <= f (dicho_up x y (fun z : R => cond_positivity (f z)) n) -> cond_positivity (f (dicho_up x y (fun z : R => cond_positivity (f z)) n)) = true
H11:cond_positivity (f y) = true -> 0 <= f y
H12:0 <= f y -> cond_positivity (f y) = true

0 <= f y
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
forall z : R, cond_positivity z = true <-> 0 <= z
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat

forall z : R, cond_positivity z = true <-> 0 <= z
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
z:R

cond_positivity z = true <-> 0 <= z
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
z:R

(if Rle_dec 0 z then true else false) = true <-> 0 <= z
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
z:R
Hle:0 <= z

true = true <-> 0 <= z
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
z:R
Hnle:~ 0 <= z
false = true <-> 0 <= z
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
z:R
Hle:0 <= z

true = true -> 0 <= z
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
z:R
Hle:0 <= z
0 <= z -> true = true
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
z:R
Hnle:~ 0 <= z
false = true <-> 0 <= z
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
z:R
Hle:0 <= z

0 <= z -> true = true
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
z:R
Hnle:~ 0 <= z
false = true <-> 0 <= z
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
z:R
Hnle:~ 0 <= z

false = true <-> 0 <= z
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
z:R
Hnle:~ 0 <= z

false = true -> 0 <= z
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
z:R
Hnle:~ 0 <= z
0 <= z -> false = true
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
z:R
Hnle:~ 0 <= z

0 <= z -> false = true
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
n:nat
z:R
Hnle:~ 0 <= z
H7:0 <= z

false = true
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0

forall n : nat, f (Vn n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0

forall n : nat, f (dicho_lb x y (fun z : R => cond_positivity (f z)) n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0

(forall z : R, cond_positivity z = false <-> z < 0) -> forall n : nat, f (dicho_lb x y (fun z : R => cond_positivity (f z)) n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall z : R, cond_positivity z = false <-> z < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
H7:forall z : R, cond_positivity z = false <-> z < 0
n:nat

f (dicho_lb x y (fun z : R => cond_positivity (f z)) n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall z : R, cond_positivity z = false <-> z < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
H7:forall z : R, cond_positivity z = false <-> z < 0
n:nat
H8:(fun z : R => cond_positivity (f z)) x = false -> (fun z : R => cond_positivity (f z)) (dicho_lb x y (fun z : R => cond_positivity (f z)) n) = false

f (dicho_lb x y (fun z : R => cond_positivity (f z)) n) <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall z : R, cond_positivity z = false <-> z < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
H7:forall z : R, cond_positivity z = false <-> z < 0
n:nat
H8:(fun z : R => cond_positivity (f z)) x = false -> (fun z : R => cond_positivity (f z)) (dicho_lb x y (fun z : R => cond_positivity (f z)) n) = false

f (dicho_lb x y (fun z : R => cond_positivity (f z)) n) < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall z : R, cond_positivity z = false <-> z < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
H7:forall z : R, cond_positivity z = false <-> z < 0
n:nat
H8:(fun z : R => cond_positivity (f z)) x = false -> (fun z : R => cond_positivity (f z)) (dicho_lb x y (fun z : R => cond_positivity (f z)) n) = false
H9:cond_positivity (f (dicho_lb x y (fun z : R => cond_positivity (f z)) n)) = false -> f (dicho_lb x y (fun z : R => cond_positivity (f z)) n) < 0
H10:f (dicho_lb x y (fun z : R => cond_positivity (f z)) n) < 0 -> cond_positivity (f (dicho_lb x y (fun z : R => cond_positivity (f z)) n)) = false

f (dicho_lb x y (fun z : R => cond_positivity (f z)) n) < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall z : R, cond_positivity z = false <-> z < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
H7:forall z : R, cond_positivity z = false <-> z < 0
n:nat
H8:(fun z : R => cond_positivity (f z)) x = false -> (fun z : R => cond_positivity (f z)) (dicho_lb x y (fun z : R => cond_positivity (f z)) n) = false
H9:cond_positivity (f (dicho_lb x y (fun z : R => cond_positivity (f z)) n)) = false -> f (dicho_lb x y (fun z : R => cond_positivity (f z)) n) < 0
H10:f (dicho_lb x y (fun z : R => cond_positivity (f z)) n) < 0 -> cond_positivity (f (dicho_lb x y (fun z : R => cond_positivity (f z)) n)) = false

cond_positivity (f (dicho_lb x y (fun z : R => cond_positivity (f z)) n)) = false
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall z : R, cond_positivity z = false <-> z < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
H7:forall z : R, cond_positivity z = false <-> z < 0
n:nat
H8:(fun z : R => cond_positivity (f z)) x = false -> (fun z : R => cond_positivity (f z)) (dicho_lb x y (fun z : R => cond_positivity (f z)) n) = false
H9:cond_positivity (f (dicho_lb x y (fun z : R => cond_positivity (f z)) n)) = false -> f (dicho_lb x y (fun z : R => cond_positivity (f z)) n) < 0
H10:f (dicho_lb x y (fun z : R => cond_positivity (f z)) n) < 0 -> cond_positivity (f (dicho_lb x y (fun z : R => cond_positivity (f z)) n)) = false

cond_positivity (f x) = false
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall z : R, cond_positivity z = false <-> z < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
H7:forall z : R, cond_positivity z = false <-> z < 0
n:nat
H8:(fun z : R => cond_positivity (f z)) x = false -> (fun z : R => cond_positivity (f z)) (dicho_lb x y (fun z : R => cond_positivity (f z)) n) = false
H9:cond_positivity (f (dicho_lb x y (fun z : R => cond_positivity (f z)) n)) = false -> f (dicho_lb x y (fun z : R => cond_positivity (f z)) n) < 0
H10:f (dicho_lb x y (fun z : R => cond_positivity (f z)) n) < 0 -> cond_positivity (f (dicho_lb x y (fun z : R => cond_positivity (f z)) n)) = false
H11:cond_positivity (f x) = false -> f x < 0
H12:f x < 0 -> cond_positivity (f x) = false

cond_positivity (f x) = false
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall z : R, cond_positivity z = false <-> z < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n0 : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n0:nat -> R
Wn:=fun n0 : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n0:nat -> R
H5:(forall n0 : nat, 0 <= f (Wn n0)) -> 0 <= f x0
H6:(forall n0 : nat, f (Vn n0) <= 0) -> f x0 <= 0
H7:forall z : R, cond_positivity z = false <-> z < 0
n:nat
H8:(fun z : R => cond_positivity (f z)) x = false -> (fun z : R => cond_positivity (f z)) (dicho_lb x y (fun z : R => cond_positivity (f z)) n) = false
H9:cond_positivity (f (dicho_lb x y (fun z : R => cond_positivity (f z)) n)) = false -> f (dicho_lb x y (fun z : R => cond_positivity (f z)) n) < 0
H10:f (dicho_lb x y (fun z : R => cond_positivity (f z)) n) < 0 -> cond_positivity (f (dicho_lb x y (fun z : R => cond_positivity (f z)) n)) = false
H11:cond_positivity (f x) = false -> f x < 0
H12:f x < 0 -> cond_positivity (f x) = false

f x < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
forall z : R, cond_positivity z = false <-> z < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0

forall z : R, cond_positivity z = false <-> z < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
z:R

cond_positivity z = false <-> z < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
z:R

(if Rle_dec 0 z then true else false) = false <-> z < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
z:R
Hle:0 <= z

true = false <-> z < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
z:R
Hnle:~ 0 <= z
false = false <-> z < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
z:R
Hle:0 <= z

true = false -> z < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
z:R
Hle:0 <= z
z < 0 -> true = false
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
z:R
Hnle:~ 0 <= z
false = false <-> z < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
z:R
Hle:0 <= z

z < 0 -> true = false
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
z:R
Hnle:~ 0 <= z
false = false <-> z < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
z:R
Hnle:~ 0 <= z

false = false <-> z < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
z:R
Hnle:~ 0 <= z

false = false -> z < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
z:R
Hnle:~ 0 <= z
z < 0 -> false = false
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z0 : R => cond_positivity (f z0))) x0
p:Un_cv (dicho_up x y (fun z0 : R => cond_positivity (f z0))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z0 : R => cond_positivity (f z0)) n:nat -> R
H5:(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
H6:(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
z:R
Hnle:~ 0 <= z

z < 0 -> false = false
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R

(forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R

Un_cv Wn x0 -> (forall n : nat, 0 <= f (Wn n)) -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)

0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:Un_cv (fun i : nat => f (Wn i)) (f x0)

0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:Un_cv (fun i : nat => f (Wn i)) (f x0)
Hlt:0 < f x0

0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:Un_cv (fun i : nat => f (Wn i)) (f x0)
0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:Un_cv (fun i : nat => f (Wn i)) (f x0)
Hgt:0 > f x0
0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:Un_cv (fun i : nat => f (Wn i)) (f x0)

0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:Un_cv (fun i : nat => f (Wn i)) (f x0)
Hgt:0 > f x0
0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:Un_cv (fun i : nat => f (Wn i)) (f x0)
Hgt:0 > f x0

0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0

0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0

0 < - f x0 -> 0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
0 < - f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
H8:0 < - f x0

0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
0 < - f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
H8:0 < - f x0
x2:nat
H9:forall n : nat, (n >= x2)%nat -> Rabs (f (Wn n) - f x0) < - f x0

0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
0 < - f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
H8:0 < - f x0
x2:nat
H9:forall n : nat, (n >= x2)%nat -> Rabs (f (Wn n) - f x0) < - f x0
H10:(x2 >= x2)%nat

0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
0 < - f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
H8:0 < - f x0
x2:nat
H9:forall n : nat, (n >= x2)%nat -> Rabs (f (Wn n) - f x0) < - f x0
H10:(x2 >= x2)%nat
H11:Rabs (f (Wn x2) - f x0) < - f x0

0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
0 < - f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
H8:0 < - f x0
x2:nat
H9:forall n : nat, (n >= x2)%nat -> Rabs (f (Wn n) - f x0) < - f x0
H10:(x2 >= x2)%nat
H11:f (Wn x2) - f x0 < - f x0

0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
H8:0 < - f x0
x2:nat
H9:forall n : nat, (n >= x2)%nat -> Rabs (f (Wn n) - f x0) < - f x0
H10:(x2 >= x2)%nat
H11:Rabs (f (Wn x2) - f x0) < - f x0
f (Wn x2) - f x0 >= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
0 < - f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
H8:0 < - f x0
x2:nat
H9:forall n : nat, (n >= x2)%nat -> Rabs (f (Wn n) - f x0) < - f x0
H10:(x2 >= x2)%nat
H11:f (Wn x2) - f x0 < - f x0 + 0

0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
H8:0 < - f x0
x2:nat
H9:forall n : nat, (n >= x2)%nat -> Rabs (f (Wn n) - f x0) < - f x0
H10:(x2 >= x2)%nat
H11:Rabs (f (Wn x2) - f x0) < - f x0
f (Wn x2) - f x0 >= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
0 < - f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
H8:0 < - f x0
x2:nat
H9:forall n : nat, (n >= x2)%nat -> Rabs (f (Wn n) - f x0) < - f x0
H10:(x2 >= x2)%nat
H11:- f x0 + f (Wn x2) < - f x0 + 0

0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
H8:0 < - f x0
x2:nat
H9:forall n : nat, (n >= x2)%nat -> Rabs (f (Wn n) - f x0) < - f x0
H10:(x2 >= x2)%nat
H11:Rabs (f (Wn x2) - f x0) < - f x0
f (Wn x2) - f x0 >= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
0 < - f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
H8:0 < - f x0
x2:nat
H9:forall n : nat, (n >= x2)%nat -> Rabs (f (Wn n) - f x0) < - f x0
H10:(x2 >= x2)%nat
H11:- f x0 + f (Wn x2) < - f x0 + 0
H12:f (Wn x2) < 0

0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
H8:0 < - f x0
x2:nat
H9:forall n : nat, (n >= x2)%nat -> Rabs (f (Wn n) - f x0) < - f x0
H10:(x2 >= x2)%nat
H11:Rabs (f (Wn x2) - f x0) < - f x0
f (Wn x2) - f x0 >= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
0 < - f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
H8:0 < - f x0
x2:nat
H9:forall n : nat, (n >= x2)%nat -> Rabs (f (Wn n) - f x0) < - f x0
H10:(x2 >= x2)%nat
H11:- f x0 + f (Wn x2) < - f x0 + 0
H12:f (Wn x2) < 0
H13:0 <= f (Wn x2)

0 <= f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
H8:0 < - f x0
x2:nat
H9:forall n : nat, (n >= x2)%nat -> Rabs (f (Wn n) - f x0) < - f x0
H10:(x2 >= x2)%nat
H11:Rabs (f (Wn x2) - f x0) < - f x0
f (Wn x2) - f x0 >= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
0 < - f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
H8:0 < - f x0
x2:nat
H9:forall n : nat, (n >= x2)%nat -> Rabs (f (Wn n) - f x0) < - f x0
H10:(x2 >= x2)%nat
H11:Rabs (f (Wn x2) - f x0) < - f x0

f (Wn x2) - f x0 >= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
0 < - f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
H8:0 < - f x0
x2:nat
H9:forall n : nat, (n >= x2)%nat -> Rabs (f (Wn n) - f x0) < - f x0
H10:(x2 >= x2)%nat
H11:Rabs (f (Wn x2) - f x0) < - f x0

0 <= f (Wn x2)
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
H8:0 < - f x0
x2:nat
H9:forall n : nat, (n >= x2)%nat -> Rabs (f (Wn n) - f x0) < - f x0
H10:(x2 >= x2)%nat
H11:Rabs (f (Wn x2) - f x0) < - f x0
0 < - f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
0 < - f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
H8:0 < - f x0
x2:nat
H9:forall n : nat, (n >= x2)%nat -> Rabs (f (Wn n) - f x0) < - f x0
H10:(x2 >= x2)%nat
H11:Rabs (f (Wn x2) - f x0) < - f x0

0 < - f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0
0 < - f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Wn x0
H6:forall n : nat, 0 <= f (Wn n)
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Wn n) - f x0) < eps
Hgt:0 > f x0

0 < - f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R

Un_cv Wn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R

(forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R

Un_cv Vn x0 -> (forall n : nat, f (Vn n) <= 0) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0

f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)

f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hlt:0 < f x0

f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hgt:0 > f x0
f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0

f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hgt:0 > f x0
f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0

f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hgt:0 > f x0
f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat

f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hgt:0 > f x0
f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:Rabs (f (Vn x2) - f x0) < f x0

f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hgt:0 > f x0
f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:- (f (Vn x2) - f x0) < f x0

f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:Rabs (f (Vn x2) - f x0) < f x0
f (Vn x2) - f x0 < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hgt:0 > f x0
f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:- (f (Vn x2) - f x0) < f x0 + 0

f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:Rabs (f (Vn x2) - f x0) < f x0
f (Vn x2) - f x0 < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hgt:0 > f x0
f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:f x0 - f (Vn x2) < f x0 + 0

f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:Rabs (f (Vn x2) - f x0) < f x0
f (Vn x2) - f x0 < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hgt:0 > f x0
f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:f x0 + - f (Vn x2) < f x0 + 0

f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:Rabs (f (Vn x2) - f x0) < f x0
f (Vn x2) - f x0 < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hgt:0 > f x0
f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:f x0 + - f (Vn x2) < f x0 + 0
H11:- f (Vn x2) < 0

f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:Rabs (f (Vn x2) - f x0) < f x0
f (Vn x2) - f x0 < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hgt:0 > f x0
f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:f x0 + - f (Vn x2) < f x0 + 0
H11:- f (Vn x2) < 0
H12:f (Vn x2) <= 0

f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:Rabs (f (Vn x2) - f x0) < f x0
f (Vn x2) - f x0 < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hgt:0 > f x0
f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:f x0 + - f (Vn x2) < f x0 + 0
H11:- f (Vn x2) < 0
H12:f (Vn x2) <= 0

0 < f (Vn x2) -> f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:f x0 + - f (Vn x2) < f x0 + 0
H11:- f (Vn x2) < 0
H12:f (Vn x2) <= 0
0 < f (Vn x2)
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:Rabs (f (Vn x2) - f x0) < f x0
f (Vn x2) - f x0 < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hgt:0 > f x0
f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:f x0 + - f (Vn x2) < f x0 + 0
H11:- f (Vn x2) < 0
H12:f (Vn x2) <= 0
H13:0 < f (Vn x2)

f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:f x0 + - f (Vn x2) < f x0 + 0
H11:- f (Vn x2) < 0
H12:f (Vn x2) <= 0
0 < f (Vn x2)
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:Rabs (f (Vn x2) - f x0) < f x0
f (Vn x2) - f x0 < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hgt:0 > f x0
f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:f x0 + - f (Vn x2) < f x0 + 0
H11:- f (Vn x2) < 0
H12:f (Vn x2) <= 0

0 < f (Vn x2)
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:Rabs (f (Vn x2) - f x0) < f x0
f (Vn x2) - f x0 < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hgt:0 > f x0
f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:f x0 + - f (Vn x2) < f x0 + 0
H11:- f (Vn x2) < 0
H12:f (Vn x2) <= 0

0 < - - f (Vn x2)
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:Rabs (f (Vn x2) - f x0) < f x0
f (Vn x2) - f x0 < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hgt:0 > f x0
f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:Rabs (f (Vn x2) - f x0) < f x0

f (Vn x2) - f x0 < 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hgt:0 > f x0
f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:Rabs (f (Vn x2) - f x0) < f x0

f x0 - f (Vn x2) + (f (Vn x2) - f x0) < f x0 - f (Vn x2) + 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hgt:0 > f x0
f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:Rabs (f (Vn x2) - f x0) < f x0

0 < f x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:Rabs (f (Vn x2) - f x0) < f x0
0 <= - f (Vn x2)
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hgt:0 > f x0
f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (f (Vn n) - f x0) < eps
Hlt:0 < f x0
x2:nat
H8:forall n : nat, (n >= x2)%nat -> Rabs (f (Vn n) - f x0) < f x0
H9:(x2 >= x2)%nat
H10:Rabs (f (Vn x2) - f x0) < f x0

0 <= - f (Vn x2)
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hgt:0 > f x0
f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)

0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hgt:0 > f x0
f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
H5:Un_cv Vn x0
H6:forall n : nat, f (Vn n) <= 0
H7:Un_cv (fun i : nat => f (Vn i)) (f x0)
Hgt:0 > f x0

f x0 <= 0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R
Un_cv Vn x0
f:R -> R
x, y:R
H:continuity f
H0:x < y
H1:f x < 0
H2:0 < f y
H3:x <= y
x1, x0:R
p0:Un_cv (dicho_lb x y (fun z : R => cond_positivity (f z))) x0
p:Un_cv (dicho_up x y (fun z : R => cond_positivity (f z))) x0
H4:x1 = x0
Vn:=fun n : nat => dicho_lb x y (fun z : R => cond_positivity (f z)) n:nat -> R
Wn:=fun n : nat => dicho_up x y (fun z : R => cond_positivity (f z)) n:nat -> R

Un_cv Vn x0
unfold Vn; assumption. Qed.

forall (f : R -> R) (x y : R), continuity f -> x <= y -> f x * f y <= 0 -> {z : R | x <= z <= y /\ f z = 0}

forall (f : R -> R) (x y : R), continuity f -> x <= y -> f x * f y <= 0 -> {z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hlty:0 < f y

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Heqy:0 = f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Heqy:0 = f y

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Heqy:0 = f y

x <= y <= y /\ f y = 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Heqy:0 = f y

x <= y <= y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Heqy:0 = f y
f y = 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Heqy:0 = f y

f y = 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y

x < y -> {z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}

(- f)%F x < 0 -> {z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
(- f)%F x < 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}

0 < (- f)%F y -> (- f)%F x < 0 -> {z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
0 < (- f)%F y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
(- f)%F x < 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
H4:0 < (- f)%F y
H5:(- f)%F x < 0

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
0 < (- f)%F y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
(- f)%F x < 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
H4:0 < (- f)%F y
H5:(- f)%F x < 0
x0:R
H6:x <= x0 <= y
H7:(- f)%F x0 = 0

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
0 < (- f)%F y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
(- f)%F x < 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
H4:0 < (- f)%F y
H5:(- f)%F x < 0
x0:R
H6:x <= x0 <= y
H7:(- f)%F x0 = 0

x <= x0 <= y /\ f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
0 < (- f)%F y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
(- f)%F x < 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
H4:0 < (- f)%F y
H5:(- f)%F x < 0
x0:R
H6:x <= x0 <= y
H7:(- f)%F x0 = 0

x <= x0 <= y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
H4:0 < (- f)%F y
H5:(- f)%F x < 0
x0:R
H6:x <= x0 <= y
H7:(- f)%F x0 = 0
f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
0 < (- f)%F y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
(- f)%F x < 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
H4:0 < (- f)%F y
H5:(- f)%F x < 0
x0:R
H6:x <= x0 <= y
H7:(- f)%F x0 = 0

f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
0 < (- f)%F y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
(- f)%F x < 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
H4:0 < (- f)%F y
H5:(- f)%F x < 0
x0:R
H6:x <= x0 <= y
H7:- f x0 = 0

f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
0 < (- f)%F y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
(- f)%F x < 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
H4:0 < (- f)%F y
H5:(- f)%F x < 0
x0:R
H6:x <= x0 <= y
H7:- f x0 = 0

- - f x0 = 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
0 < (- f)%F y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
(- f)%F x < 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}

0 < (- f)%F y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}
(- f)%F x < 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}

(- f)%F x < 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y
H3:(- f)%F x < 0 -> 0 < (- f)%F y -> {z : R | x <= z <= y /\ (- f)%F z = 0}

- f x < 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y

x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x < y

x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x = y
x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f x
Hgty:0 > f y
H2:x = y

x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hltx:0 < f y
Hgty:0 > f y
H2:x = y

x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x

x <= x <= y /\ f x = 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x

x <= x <= y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x
f x = 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Heqx:0 = f x

f x = 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hlty:0 < f y

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Heqy:0 = f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hgty:0 > f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hlty:0 < f y

x < y -> {z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hlty:0 < f y
x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Heqy:0 = f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hgty:0 > f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hlty:0 < f y
H2:x < y

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hlty:0 < f y
x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Heqy:0 = f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hgty:0 > f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hlty:0 < f y

x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Heqy:0 = f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hgty:0 > f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hlty:0 < f y
H2:x < y

x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hlty:0 < f y
H2:x = y
x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Heqy:0 = f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hgty:0 > f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hlty:0 < f y
H2:x = y

x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Heqy:0 = f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hgty:0 > f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f y
Hlty:0 < f y
H2:x = y

x < y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Heqy:0 = f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hgty:0 > f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Heqy:0 = f y

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hgty:0 > f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Heqy:0 = f y

x <= y <= y /\ f y = 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hgty:0 > f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Heqy:0 = f y

x <= y <= y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Heqy:0 = f y
f y = 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hgty:0 > f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Heqy:0 = f y

f y = 0
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hgty:0 > f y
{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hgty:0 > f y

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hgty:0 > f y

0 < f x * f y -> {z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hgty:0 > f y
0 < f x * f y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hgty:0 > f y
H2:0 < f x * f y

{z : R | x <= z <= y /\ f z = 0}
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hgty:0 > f y
0 < f x * f y
f:R -> R
x, y:R
H:continuity f
H0:x <= y
H1:f x * f y <= 0
Hgtx:0 > f x
Hgty:0 > f y

0 < f x * f y
rewrite <- Rmult_opp_opp; apply Rmult_lt_0_compat; apply Ropp_0_gt_lt_contravar; assumption. Qed.
We can now define the square root function as the reciprocal transformation of the square function

forall y : R, 0 <= y -> {z : R | 0 <= z /\ y = z²}

forall y : R, 0 <= y -> {z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y

{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R

{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R

f 0 <= 0 -> {z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0

{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0

continuity f -> {z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f

{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1

{z : R | 0 <= z /\ y = z²}
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1
{z : R | 0 <= z /\ 1 = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1

0 <= f 1 -> {z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
0 <= f 1
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1
{z : R | 0 <= z /\ 1 = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1

{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
0 <= f 1
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1
{z : R | 0 <= z /\ 1 = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1

f 0 * f 1 <= 0 -> {z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1
f 0 * f 1 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
0 <= f 1
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1
{z : R | 0 <= z /\ 1 = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1
H3:f 0 * f 1 <= 0

{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1
f 0 * f 1 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
0 <= f 1
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1
{z : R | 0 <= z /\ 1 = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1
H3:f 0 * f 1 <= 0
X:{z : R | 0 <= z <= 1 /\ f z = 0}

{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1
f 0 * f 1 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
0 <= f 1
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1
{z : R | 0 <= z /\ 1 = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1
H3:f 0 * f 1 <= 0
X:{z : R | 0 <= z <= 1 /\ f z = 0}
t:R
H4:0 <= t <= 1 /\ f t = 0

{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1
f 0 * f 1 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
0 <= f 1
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1
{z : R | 0 <= z /\ 1 = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1
H3:f 0 * f 1 <= 0
X:{z : R | 0 <= z <= 1 /\ f z = 0}
t:R
H4:0 <= t <= 1 /\ f t = 0

0 <= t /\ y = t²
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1
f 0 * f 1 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
0 <= f 1
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1
{z : R | 0 <= z /\ 1 = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1
H3:f 0 * f 1 <= 0
X:{z : R | 0 <= z <= 1 /\ f z = 0}
t:R
H4:0 <= t <= 1 /\ f t = 0
H5:0 <= t <= 1
H6:f t = 0

0 <= t /\ y = t²
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1
f 0 * f 1 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
0 <= f 1
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1
{z : R | 0 <= z /\ 1 = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1
H3:f 0 * f 1 <= 0
X:{z : R | 0 <= z <= 1 /\ f z = 0}
t:R
H4:0 <= t <= 1 /\ f t = 0
H5:0 <= t <= 1
H6:f t = 0

0 <= t
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1
H3:f 0 * f 1 <= 0
X:{z : R | 0 <= z <= 1 /\ f z = 0}
t:R
H4:0 <= t <= 1 /\ f t = 0
H5:0 <= t <= 1
H6:f t = 0
y = t²
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1
f 0 * f 1 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
0 <= f 1
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1
{z : R | 0 <= z /\ 1 = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1
H3:f 0 * f 1 <= 0
X:{z : R | 0 <= z <= 1 /\ f z = 0}
t:R
H4:0 <= t <= 1 /\ f t = 0
H5:0 <= t <= 1
H6:f t = 0

y = t²
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1
f 0 * f 1 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
0 <= f 1
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1
{z : R | 0 <= z /\ 1 = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1
H3:f 0 * f 1 <= 0
X:{z : R | 0 <= z <= 1 /\ f z = 0}
t:R
H4:0 <= t <= 1 /\ f t = 0
H5:0 <= t <= 1
H6:t² - y = 0

y = t²
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1
f 0 * f 1 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
0 <= f 1
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1
{z : R | 0 <= z /\ 1 = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1

f 0 * f 1 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
0 <= f 1
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1
{z : R | 0 <= z /\ 1 = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
H2:0 <= f 1

f 1 * f 0 <= f 1 * 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1
0 <= f 1
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1
{z : R | 0 <= z /\ 1 = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1

0 <= f 1
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1
{z : R | 0 <= z /\ 1 = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1

0 <= 1² - y
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1
{z : R | 0 <= z /\ 1 = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1

0 <= 1 - y
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1
{z : R | 0 <= z /\ 1 = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hlt:y < 1

y + 0 <= y + (1 - y)
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1
{z : R | 0 <= z /\ 1 = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1

{z : R | 0 <= z /\ 1 = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1

0 <= 1 /\ 1 = 1²
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1

0 <= 1
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1
1 = 1²
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
f:=fun x : R => x² - 1:R -> R
H1:continuity f
H0:f 0 <= 0
H:0 <= 1

1 = 1²
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1

{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1

0 <= f y -> {z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
0 <= f y
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y

{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
0 <= f y
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y

f 0 * f y <= 0 -> {z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y
f 0 * f y <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
0 <= f y
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y
H3:f 0 * f y <= 0

{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y
f 0 * f y <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
0 <= f y
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y
H3:f 0 * f y <= 0
X:{z : R | 0 <= z <= y /\ f z = 0}

{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y
f 0 * f y <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
0 <= f y
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y
H3:f 0 * f y <= 0
X:{z : R | 0 <= z <= y /\ f z = 0}
t:R
H4:0 <= t <= y /\ f t = 0

{z : R | 0 <= z /\ y = z²}
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y
f 0 * f y <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
0 <= f y
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y
H3:f 0 * f y <= 0
X:{z : R | 0 <= z <= y /\ f z = 0}
t:R
H4:0 <= t <= y /\ f t = 0

0 <= t /\ y = t²
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y
f 0 * f y <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
0 <= f y
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y
H3:f 0 * f y <= 0
X:{z : R | 0 <= z <= y /\ f z = 0}
t:R
H4:0 <= t <= y /\ f t = 0
H5:0 <= t <= y
H6:f t = 0

0 <= t /\ y = t²
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y
f 0 * f y <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
0 <= f y
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y
H3:f 0 * f y <= 0
X:{z : R | 0 <= z <= y /\ f z = 0}
t:R
H4:0 <= t <= y /\ f t = 0
H5:0 <= t <= y
H6:f t = 0

0 <= t
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y
H3:f 0 * f y <= 0
X:{z : R | 0 <= z <= y /\ f z = 0}
t:R
H4:0 <= t <= y /\ f t = 0
H5:0 <= t <= y
H6:f t = 0
y = t²
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y
f 0 * f y <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
0 <= f y
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y
H3:f 0 * f y <= 0
X:{z : R | 0 <= z <= y /\ f z = 0}
t:R
H4:0 <= t <= y /\ f t = 0
H5:0 <= t <= y
H6:f t = 0

y = t²
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y
f 0 * f y <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
0 <= f y
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y
H3:f 0 * f y <= 0
X:{z : R | 0 <= z <= y /\ f z = 0}
t:R
H4:0 <= t <= y /\ f t = 0
H5:0 <= t <= y
H6:t² - y = 0

y = t²
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y
f 0 * f y <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
0 <= f y
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y

f 0 * f y <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
0 <= f y
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
H2:0 <= f y

f y * f 0 <= f y * 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
0 <= f y
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1

0 <= f y
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1

0 <= y² - y
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1

y + 0 <= y + (y² - y)
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1

y <= y²
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1

y * 1 <= y²
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1

0 <= y
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1
1 <= y
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
H1:continuity f
Hgt:y > 1

1 <= y
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0

continuity f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0

continuity (Rsqr - fct_cte y)
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
(Rsqr - fct_cte y)%F = f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0

continuity Rsqr
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
continuity (fct_cte y)
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
(Rsqr - fct_cte y)%F = f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0

continuity (fct_cte y)
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0
(Rsqr - fct_cte y)%F = f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
H0:f 0 <= 0

(Rsqr - fct_cte y)%F = f
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R
f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R

f 0 <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R

0 - y <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R

- y <= 0
y:R
H:0 <= y
f:=fun x : R => x² - y:R -> R

0 >= - y
apply Ropp_0_le_ge_contravar; assumption. Qed. (* Definition of the square root: R+->R *) Definition Rsqrt (y:nonnegreal) : R := let (a,_) := Rsqrt_exists (nonneg y) (cond_nonneg y) in a. (**********)

forall x : nonnegreal, 0 <= Rsqrt x

forall x : nonnegreal, 0 <= Rsqrt x
x:nonnegreal

0 <= Rsqrt x
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²

0 <= Rsqrt x
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²

x0 = Rsqrt x -> 0 <= Rsqrt x
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
x0 = Rsqrt x
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
H:x0 = Rsqrt x

0 <= Rsqrt x
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
x0 = Rsqrt x
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²

x0 = Rsqrt x
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²

x0 = (let (a, _) := Rsqrt_exists x (cond_nonneg x) in a)
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
x1:R
H:0 <= x1
H0:x = x1²

x0 = x1
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
x1:R
H:0 <= x1
H0:x = x1²

0 <= x0
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
x1:R
H:0 <= x1
H0:x = x1²
0 <= x1
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
x1:R
H:0 <= x1
H0:x = x1²
x0² = x1²
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
x1:R
H:0 <= x1
H0:x = x1²

0 <= x1
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
x1:R
H:0 <= x1
H0:x = x1²
x0² = x1²
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
x1:R
H:0 <= x1
H0:x = x1²

x0² = x1²
rewrite <- H0, <- H2; reflexivity. Qed. (**********)

forall x : nonnegreal, Rsqrt x * Rsqrt x = x

forall x : nonnegreal, Rsqrt x * Rsqrt x = x
x:nonnegreal

Rsqrt x * Rsqrt x = x
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²

Rsqrt x * Rsqrt x = x
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²

x0 = Rsqrt x -> Rsqrt x * Rsqrt x = x
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
x0 = Rsqrt x
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
H:x0 = Rsqrt x

Rsqrt x * Rsqrt x = x
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
x0 = Rsqrt x
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
H:x0 = Rsqrt x

x0 * x0 = x
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
x0 = Rsqrt x
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²

x0 = Rsqrt x
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²

x0 = (let (a, _) := Rsqrt_exists x (cond_nonneg x) in a)
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
x1:R
H:0 <= x1
H0:x = x1²

x0 = x1
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
x1:R
H:0 <= x1
H0:x = x1²

0 <= x0
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
x1:R
H:0 <= x1
H0:x = x1²
0 <= x1
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
x1:R
H:0 <= x1
H0:x = x1²
x0² = x1²
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
x1:R
H:0 <= x1
H0:x = x1²

0 <= x1
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
x1:R
H:0 <= x1
H0:x = x1²
x0² = x1²
x:nonnegreal
x0:R
H1:0 <= x0
H2:x = x0²
x1:R
H:0 <= x1
H0:x = x1²

x0² = x1²
rewrite <- H0, <- H2; reflexivity. Qed.