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

Require Import Rbase.
Require Import Rfunctions.
Require Import Ranalysis1.
Require Import RList.
Require Import Classical_Prop.
Require Import Classical_Pred_Type.
Local Open Scope R_scope.

General definitions and propositions

Definition included (D1 D2:R -> Prop) : Prop := forall x:R, D1 x -> D2 x.
Definition disc (x:R) (delta:posreal) (y:R) : Prop := Rabs (y - x) < delta.
Definition neighbourhood (V:R -> Prop) (x:R) : Prop :=
  exists delta : posreal, included (disc x delta) V.
Definition open_set (D:R -> Prop) : Prop :=
  forall x:R, D x -> neighbourhood D x.
Definition complementary (D:R -> Prop) (c:R) : Prop := ~ D c.
Definition closed_set (D:R -> Prop) : Prop := open_set (complementary D).
Definition intersection_domain (D1 D2:R -> Prop) (c:R) : Prop := D1 c /\ D2 c.
Definition union_domain (D1 D2:R -> Prop) (c:R) : Prop := D1 c \/ D2 c.
Definition interior (D:R -> Prop) (x:R) : Prop := neighbourhood D x.


forall D : R -> Prop, included (interior D) D

forall D : R -> Prop, included (interior D) D
intros; unfold included; unfold interior; intros; unfold neighbourhood in H; elim H; intros; unfold included in H0; apply H0; unfold disc; unfold Rminus; rewrite Rplus_opp_r; rewrite Rabs_R0; apply (cond_pos x0). Qed.

forall D : R -> Prop, open_set D -> included D (interior D)

forall D : R -> Prop, open_set D -> included D (interior D)
intros; unfold open_set in H; unfold included; intros; assert (H1 := H _ H0); unfold interior; apply H1. Qed. Definition point_adherent (D:R -> Prop) (x:R) : Prop := forall V:R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V D y. Definition adherence (D:R -> Prop) (x:R) : Prop := point_adherent D x.

forall D : R -> Prop, included D (adherence D)

forall D : R -> Prop, included D (adherence D)
D:R -> Prop
x:R
H:D x
V:R -> Prop
H0:neighbourhood V x

V x
D:R -> Prop
x:R
H:D x
V:R -> Prop
H0:neighbourhood V x
D x
D:R -> Prop
x:R
H:D x
V:R -> Prop
H0:neighbourhood V x

D x
apply H. Qed.

forall D1 D2 D3 : R -> Prop, included D1 D2 -> included D2 D3 -> included D1 D3

forall D1 D2 D3 : R -> Prop, included D1 D2 -> included D2 D3 -> included D1 D3
unfold included; intros; apply H0; apply H; apply H1. Qed.

forall D : R -> Prop, open_set (interior D)

forall D : R -> Prop, open_set (interior D)
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D

exists delta : posreal, included (disc x delta) (fun x1 : R => exists delta0 : posreal, included (disc x1 delta0) D)
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1

exists delta : posreal, forall x2 : R, disc x1 delta x2 -> D x2
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R

exists delta : posreal, forall x2 : R, disc x1 delta x2 -> D x2
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R

0 < del -> exists delta : posreal, forall x2 : R, disc x1 delta x2 -> D x2
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
0 < del
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
H2:0 < del
x2:R
H3:disc x1 {| pos := del; cond_pos := H2 |} x2

D x2
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
0 < del
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
H2:0 < del
x2:R
H3:disc x1 {| pos := del; cond_pos := H2 |} x2

included (disc x1 {| pos := del; cond_pos := H2 |}) (disc x x0) -> D x2
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
H2:0 < del
x2:R
H3:disc x1 {| pos := del; cond_pos := H2 |} x2
included (disc x1 {| pos := del; cond_pos := H2 |}) (disc x x0)
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
0 < del
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
H2:0 < del
x2:R
H3:disc x1 {| pos := del; cond_pos := H2 |} x2
H4:included (disc x1 {| pos := del; cond_pos := H2 |}) (disc x x0)
H5:included (disc x1 {| pos := del; cond_pos := H2 |}) D

D x2
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
H2:0 < del
x2:R
H3:disc x1 {| pos := del; cond_pos := H2 |} x2
included (disc x1 {| pos := del; cond_pos := H2 |}) (disc x x0)
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
0 < del
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
H2:0 < del
x2:R
H3:disc x1 {| pos := del; cond_pos := H2 |} x2

included (disc x1 {| pos := del; cond_pos := H2 |}) (disc x x0)
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
0 < del
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
H2:0 < del
x2:R
H3:disc x1 {| pos := del; cond_pos := H2 |} x2
x3:R
H4:Rabs (x3 - x1) < {| pos := del; cond_pos := H2 |}

Rabs (x3 - x) < x0
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
0 < del
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
H2:0 < del
x2:R
H3:disc x1 {| pos := del; cond_pos := H2 |} x2
x3:R
H4:Rabs (x3 - x1) < {| pos := del; cond_pos := H2 |}

Rabs (x3 - x) <= Rabs (x3 - x1) + Rabs (x1 - x)
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
H2:0 < del
x2:R
H3:disc x1 {| pos := del; cond_pos := H2 |} x2
x3:R
H4:Rabs (x3 - x1) < {| pos := del; cond_pos := H2 |}
Rabs (x3 - x1) + Rabs (x1 - x) < x0
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
0 < del
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
H2:0 < del
x2:R
H3:disc x1 {| pos := del; cond_pos := H2 |} x2
x3:R
H4:Rabs (x3 - x1) < {| pos := del; cond_pos := H2 |}

Rabs (x3 - x1) + Rabs (x1 - x) < x0
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
0 < del
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
H2:0 < del
x2:R
H3:disc x1 {| pos := del; cond_pos := H2 |} x2
x3:R
H4:Rabs (x3 - x1) < {| pos := del; cond_pos := H2 |}

Rabs (x3 - x1) + Rabs (x1 - x) < del + Rabs (x1 - x)
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
H2:0 < del
x2:R
H3:disc x1 {| pos := del; cond_pos := H2 |} x2
x3:R
H4:Rabs (x3 - x1) < {| pos := del; cond_pos := H2 |}
del + Rabs (x1 - x) = x0
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
0 < del
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
H2:0 < del
x2:R
H3:disc x1 {| pos := del; cond_pos := H2 |} x2
x3:R
H4:Rabs (x3 - x1) < {| pos := del; cond_pos := H2 |}

del + Rabs (x1 - x) = x0
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R
0 < del
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R

0 < del
D:R -> Prop
x:R
H:exists delta : posreal, included (disc x delta) D
x0:posreal
H0:included (disc x x0) D
x1:R
H1:disc x x0 x1
del:=x0 - Rabs (x - x1):R

Rabs (x - x1) < x0
unfold disc in H1; rewrite <- Rabs_Ropp; rewrite Ropp_minus_distr; apply H1. Qed.

forall D : R -> Prop, ~ (exists y : R, intersection_domain D (complementary D) y)

forall D : R -> Prop, ~ (exists y : R, intersection_domain D (complementary D) y)
intro; red; intro; elim H; intros; unfold intersection_domain, complementary in H0; elim H0; intros; elim H2; assumption. Qed.

forall D : R -> Prop, closed_set D -> included (adherence D) D

forall D : R -> Prop, closed_set D -> included (adherence D) D
D:R -> Prop
H:forall x0 : R, ~ D x0 -> neighbourhood (fun c : R => ~ D c) x0
x:R
H0:point_adherent D x
H1:D x \/ ~ D x
H2:D x

D x
D:R -> Prop
H:forall x0 : R, ~ D x0 -> neighbourhood (fun c : R => ~ D c) x0
x:R
H0:point_adherent D x
H1:D x \/ ~ D x
H2:~ D x
D x
D:R -> Prop
H:forall x0 : R, ~ D x0 -> neighbourhood (fun c : R => ~ D c) x0
x:R
H0:point_adherent D x
H1:D x \/ ~ D x
H2:~ D x

D x
assert (H3 := H _ H2); assert (H4 := H0 _ H3); elim H4; intros; unfold intersection_domain in H5; elim H5; intros; elim H6; assumption. Qed.

forall D : R -> Prop, closed_set (adherence D)

forall D : R -> Prop, closed_set (adherence D)
D:R -> Prop
x:R
H:~ (forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V D y)
P:=fun V : R -> Prop => neighbourhood V x -> exists y : R, intersection_domain V D y:(R -> Prop) -> Prop
H0:exists n : R -> Prop, ~ P n
V0:R -> Prop
H1:~ (neighbourhood V0 x -> exists y : R, intersection_domain V0 D y)
H2:neighbourhood V0 x /\ ~ (exists y : R, intersection_domain V0 D y)
H3:exists delta : posreal, included (disc x delta) V0
H4:~ (exists y : R, intersection_domain V0 D y)
x0:posreal
H5:included (disc x x0) V0
x1:R
H6:disc x x0 x1
H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D y

False
D:R -> Prop
x:R
H:~ (forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V D y)
P:=fun V : R -> Prop => neighbourhood V x -> exists y : R, intersection_domain V D y:(R -> Prop) -> Prop
H0:exists n : R -> Prop, ~ P n
V0:R -> Prop
H1:~ (neighbourhood V0 x -> exists y : R, intersection_domain V0 D y)
H2:neighbourhood V0 x /\ ~ (exists y : R, intersection_domain V0 D y)
H3:exists delta : posreal, included (disc x delta) V0
H4:~ (exists y : R, intersection_domain V0 D y)
x0:posreal
H5:included (disc x x0) V0
x1:R
H6:disc x x0 x1
H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D y
H8:(exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2) -> exists y : R, intersection_domain V0 D y

(exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2) -> False
D:R -> Prop
x:R
H:~ (forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V D y)
P:=fun V : R -> Prop => neighbourhood V x -> exists y : R, intersection_domain V D y:(R -> Prop) -> Prop
H0:exists n : R -> Prop, ~ P n
V0:R -> Prop
H1:~ (neighbourhood V0 x -> exists y : R, intersection_domain V0 D y)
H2:neighbourhood V0 x /\ ~ (exists y : R, intersection_domain V0 D y)
H3:exists delta : posreal, included (disc x delta) V0
H4:~ (exists y : R, intersection_domain V0 D y)
x0:posreal
H5:included (disc x x0) V0
x1:R
H6:disc x x0 x1
H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D y
H8:(exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2) -> exists y : R, intersection_domain V0 D y
exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2
D:R -> Prop
x:R
H:~ (forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V D y)
P:=fun V : R -> Prop => neighbourhood V x -> exists y : R, intersection_domain V D y:(R -> Prop) -> Prop
H0:exists n : R -> Prop, ~ P n
V0:R -> Prop
H1:~ (neighbourhood V0 x -> exists y : R, intersection_domain V0 D y)
H2:neighbourhood V0 x /\ ~ (exists y : R, intersection_domain V0 D y)
H3:exists delta : posreal, included (disc x delta) V0
H4:~ (exists y : R, intersection_domain V0 D y)
x0:posreal
H5:included (disc x x0) V0
x1:R
H6:disc x x0 x1
H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D y
H8:(exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2) -> exists y : R, intersection_domain V0 D y

exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2
D:R -> Prop
x:R
H:~ (forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V D y)
P:=fun V : R -> Prop => neighbourhood V x -> exists y : R, intersection_domain V D y:(R -> Prop) -> Prop
H0:exists n : R -> Prop, ~ P n
V0:R -> Prop
H1:~ (neighbourhood V0 x -> exists y : R, intersection_domain V0 D y)
H2:neighbourhood V0 x /\ ~ (exists y : R, intersection_domain V0 D y)
H3:exists delta : posreal, included (disc x delta) V0
H4:~ (exists y : R, intersection_domain V0 D y)
x0:posreal
H5:included (disc x x0) V0
x1:R
H6:disc x x0 x1
H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D y
H8:(exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2) -> exists y : R, intersection_domain V0 D y

0 < x0 - Rabs (x - x1) -> exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2
D:R -> Prop
x:R
H:~ (forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V D y)
P:=fun V : R -> Prop => neighbourhood V x -> exists y : R, intersection_domain V D y:(R -> Prop) -> Prop
H0:exists n : R -> Prop, ~ P n
V0:R -> Prop
H1:~ (neighbourhood V0 x -> exists y : R, intersection_domain V0 D y)
H2:neighbourhood V0 x /\ ~ (exists y : R, intersection_domain V0 D y)
H3:exists delta : posreal, included (disc x delta) V0
H4:~ (exists y : R, intersection_domain V0 D y)
x0:posreal
H5:included (disc x x0) V0
x1:R
H6:disc x x0 x1
H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D y
H8:(exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2) -> exists y : R, intersection_domain V0 D y
0 < x0 - Rabs (x - x1)
D:R -> Prop
x:R
H:~ (forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V D y)
P:=fun V : R -> Prop => neighbourhood V x -> exists y : R, intersection_domain V D y:(R -> Prop) -> Prop
H0:exists n : R -> Prop, ~ P n
V0:R -> Prop
H1:~ (neighbourhood V0 x -> exists y : R, intersection_domain V0 D y)
H2:neighbourhood V0 x /\ ~ (exists y : R, intersection_domain V0 D y)
H3:exists delta : posreal, included (disc x delta) V0
H4:~ (exists y : R, intersection_domain V0 D y)
x0:posreal
H5:forall x3 : R, disc x x0 x3 -> V0 x3
x1:R
H6:disc x x0 x1
H7:forall V : R -> Prop, (exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V x3) -> exists y : R, intersection_domain V D y
H8:(exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V0 x3) -> exists y : R, intersection_domain V0 D y
H9:0 < x0 - Rabs (x - x1)
del:={| pos := x0 - Rabs (x - x1); cond_pos := H9 |}:posreal
x2:R
H10:disc x1 del x2

Rabs (x2 - x) <= Rabs (x2 - x1) + Rabs (x1 - x)
D:R -> Prop
x:R
H:~ (forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V D y)
P:=fun V : R -> Prop => neighbourhood V x -> exists y : R, intersection_domain V D y:(R -> Prop) -> Prop
H0:exists n : R -> Prop, ~ P n
V0:R -> Prop
H1:~ (neighbourhood V0 x -> exists y : R, intersection_domain V0 D y)
H2:neighbourhood V0 x /\ ~ (exists y : R, intersection_domain V0 D y)
H3:exists delta : posreal, included (disc x delta) V0
H4:~ (exists y : R, intersection_domain V0 D y)
x0:posreal
H5:forall x3 : R, disc x x0 x3 -> V0 x3
x1:R
H6:disc x x0 x1
H7:forall V : R -> Prop, (exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V x3) -> exists y : R, intersection_domain V D y
H8:(exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V0 x3) -> exists y : R, intersection_domain V0 D y
H9:0 < x0 - Rabs (x - x1)
del:={| pos := x0 - Rabs (x - x1); cond_pos := H9 |}:posreal
x2:R
H10:disc x1 del x2
Rabs (x2 - x1) + Rabs (x1 - x) < x0
D:R -> Prop
x:R
H:~ (forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V D y)
P:=fun V : R -> Prop => neighbourhood V x -> exists y : R, intersection_domain V D y:(R -> Prop) -> Prop
H0:exists n : R -> Prop, ~ P n
V0:R -> Prop
H1:~ (neighbourhood V0 x -> exists y : R, intersection_domain V0 D y)
H2:neighbourhood V0 x /\ ~ (exists y : R, intersection_domain V0 D y)
H3:exists delta : posreal, included (disc x delta) V0
H4:~ (exists y : R, intersection_domain V0 D y)
x0:posreal
H5:included (disc x x0) V0
x1:R
H6:disc x x0 x1
H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D y
H8:(exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2) -> exists y : R, intersection_domain V0 D y
0 < x0 - Rabs (x - x1)
D:R -> Prop
x:R
H:~ (forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V D y)
P:=fun V : R -> Prop => neighbourhood V x -> exists y : R, intersection_domain V D y:(R -> Prop) -> Prop
H0:exists n : R -> Prop, ~ P n
V0:R -> Prop
H1:~ (neighbourhood V0 x -> exists y : R, intersection_domain V0 D y)
H2:neighbourhood V0 x /\ ~ (exists y : R, intersection_domain V0 D y)
H3:exists delta : posreal, included (disc x delta) V0
H4:~ (exists y : R, intersection_domain V0 D y)
x0:posreal
H5:forall x3 : R, disc x x0 x3 -> V0 x3
x1:R
H6:disc x x0 x1
H7:forall V : R -> Prop, (exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V x3) -> exists y : R, intersection_domain V D y
H8:(exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V0 x3) -> exists y : R, intersection_domain V0 D y
H9:0 < x0 - Rabs (x - x1)
del:={| pos := x0 - Rabs (x - x1); cond_pos := H9 |}:posreal
x2:R
H10:disc x1 del x2

Rabs (x2 - x1) + Rabs (x1 - x) < x0
D:R -> Prop
x:R
H:~ (forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V D y)
P:=fun V : R -> Prop => neighbourhood V x -> exists y : R, intersection_domain V D y:(R -> Prop) -> Prop
H0:exists n : R -> Prop, ~ P n
V0:R -> Prop
H1:~ (neighbourhood V0 x -> exists y : R, intersection_domain V0 D y)
H2:neighbourhood V0 x /\ ~ (exists y : R, intersection_domain V0 D y)
H3:exists delta : posreal, included (disc x delta) V0
H4:~ (exists y : R, intersection_domain V0 D y)
x0:posreal
H5:included (disc x x0) V0
x1:R
H6:disc x x0 x1
H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D y
H8:(exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2) -> exists y : R, intersection_domain V0 D y
0 < x0 - Rabs (x - x1)
D:R -> Prop
x:R
H:~ (forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V D y)
P:=fun V : R -> Prop => neighbourhood V x -> exists y : R, intersection_domain V D y:(R -> Prop) -> Prop
H0:exists n : R -> Prop, ~ P n
V0:R -> Prop
H1:~ (neighbourhood V0 x -> exists y : R, intersection_domain V0 D y)
H2:neighbourhood V0 x /\ ~ (exists y : R, intersection_domain V0 D y)
H3:exists delta : posreal, included (disc x delta) V0
H4:~ (exists y : R, intersection_domain V0 D y)
x0:posreal
H5:forall x3 : R, disc x x0 x3 -> V0 x3
x1:R
H6:disc x x0 x1
H7:forall V : R -> Prop, (exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V x3) -> exists y : R, intersection_domain V D y
H8:(exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V0 x3) -> exists y : R, intersection_domain V0 D y
H9:0 < x0 - Rabs (x - x1)
del:={| pos := x0 - Rabs (x - x1); cond_pos := H9 |}:posreal
x2:R
H10:disc x1 del x2

Rabs (x2 - x1) + Rabs (x1 - x) < del + Rabs (x1 - x)
D:R -> Prop
x:R
H:~ (forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V D y)
P:=fun V : R -> Prop => neighbourhood V x -> exists y : R, intersection_domain V D y:(R -> Prop) -> Prop
H0:exists n : R -> Prop, ~ P n
V0:R -> Prop
H1:~ (neighbourhood V0 x -> exists y : R, intersection_domain V0 D y)
H2:neighbourhood V0 x /\ ~ (exists y : R, intersection_domain V0 D y)
H3:exists delta : posreal, included (disc x delta) V0
H4:~ (exists y : R, intersection_domain V0 D y)
x0:posreal
H5:forall x3 : R, disc x x0 x3 -> V0 x3
x1:R
H6:disc x x0 x1
H7:forall V : R -> Prop, (exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V x3) -> exists y : R, intersection_domain V D y
H8:(exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V0 x3) -> exists y : R, intersection_domain V0 D y
H9:0 < x0 - Rabs (x - x1)
del:={| pos := x0 - Rabs (x - x1); cond_pos := H9 |}:posreal
x2:R
H10:disc x1 del x2
del + Rabs (x1 - x) = x0
D:R -> Prop
x:R
H:~ (forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V D y)
P:=fun V : R -> Prop => neighbourhood V x -> exists y : R, intersection_domain V D y:(R -> Prop) -> Prop
H0:exists n : R -> Prop, ~ P n
V0:R -> Prop
H1:~ (neighbourhood V0 x -> exists y : R, intersection_domain V0 D y)
H2:neighbourhood V0 x /\ ~ (exists y : R, intersection_domain V0 D y)
H3:exists delta : posreal, included (disc x delta) V0
H4:~ (exists y : R, intersection_domain V0 D y)
x0:posreal
H5:included (disc x x0) V0
x1:R
H6:disc x x0 x1
H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D y
H8:(exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2) -> exists y : R, intersection_domain V0 D y
0 < x0 - Rabs (x - x1)
D:R -> Prop
x:R
H:~ (forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V D y)
P:=fun V : R -> Prop => neighbourhood V x -> exists y : R, intersection_domain V D y:(R -> Prop) -> Prop
H0:exists n : R -> Prop, ~ P n
V0:R -> Prop
H1:~ (neighbourhood V0 x -> exists y : R, intersection_domain V0 D y)
H2:neighbourhood V0 x /\ ~ (exists y : R, intersection_domain V0 D y)
H3:exists delta : posreal, included (disc x delta) V0
H4:~ (exists y : R, intersection_domain V0 D y)
x0:posreal
H5:forall x3 : R, disc x x0 x3 -> V0 x3
x1:R
H6:disc x x0 x1
H7:forall V : R -> Prop, (exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V x3) -> exists y : R, intersection_domain V D y
H8:(exists delta : posreal, forall x3 : R, disc x1 delta x3 -> V0 x3) -> exists y : R, intersection_domain V0 D y
H9:0 < x0 - Rabs (x - x1)
del:={| pos := x0 - Rabs (x - x1); cond_pos := H9 |}:posreal
x2:R
H10:disc x1 del x2

del + Rabs (x1 - x) = x0
D:R -> Prop
x:R
H:~ (forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V D y)
P:=fun V : R -> Prop => neighbourhood V x -> exists y : R, intersection_domain V D y:(R -> Prop) -> Prop
H0:exists n : R -> Prop, ~ P n
V0:R -> Prop
H1:~ (neighbourhood V0 x -> exists y : R, intersection_domain V0 D y)
H2:neighbourhood V0 x /\ ~ (exists y : R, intersection_domain V0 D y)
H3:exists delta : posreal, included (disc x delta) V0
H4:~ (exists y : R, intersection_domain V0 D y)
x0:posreal
H5:included (disc x x0) V0
x1:R
H6:disc x x0 x1
H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D y
H8:(exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2) -> exists y : R, intersection_domain V0 D y
0 < x0 - Rabs (x - x1)
D:R -> Prop
x:R
H:~ (forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V D y)
P:=fun V : R -> Prop => neighbourhood V x -> exists y : R, intersection_domain V D y:(R -> Prop) -> Prop
H0:exists n : R -> Prop, ~ P n
V0:R -> Prop
H1:~ (neighbourhood V0 x -> exists y : R, intersection_domain V0 D y)
H2:neighbourhood V0 x /\ ~ (exists y : R, intersection_domain V0 D y)
H3:exists delta : posreal, included (disc x delta) V0
H4:~ (exists y : R, intersection_domain V0 D y)
x0:posreal
H5:included (disc x x0) V0
x1:R
H6:disc x x0 x1
H7:forall V : R -> Prop, (exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V x2) -> exists y : R, intersection_domain V D y
H8:(exists delta : posreal, forall x2 : R, disc x1 delta x2 -> V0 x2) -> exists y : R, intersection_domain V0 D y

0 < x0 - Rabs (x - x1)
apply Rplus_lt_reg_l with (Rabs (x - x1)); rewrite Rplus_0_r; replace (Rabs (x - x1) + (x0 - Rabs (x - x1))) with (pos x0); [ rewrite <- Rabs_Ropp; rewrite Ropp_minus_distr; apply H6 | ring ]. Qed. Definition eq_Dom (D1 D2:R -> Prop) : Prop := included D1 D2 /\ included D2 D1. Infix "=_D" := eq_Dom (at level 70, no associativity).

forall D : R -> Prop, open_set D <-> D =_D interior D

forall D : R -> Prop, open_set D <-> D =_D interior D
D:R -> Prop

open_set D -> D =_D interior D
D:R -> Prop
D =_D interior D -> open_set D
D:R -> Prop
H:open_set D

included D (interior D)
D:R -> Prop
H:open_set D
included (interior D) D
D:R -> Prop
D =_D interior D -> open_set D
D:R -> Prop
H:open_set D

included (interior D) D
D:R -> Prop
D =_D interior D -> open_set D
D:R -> Prop

D =_D interior D -> open_set D
intro; unfold eq_Dom in H; elim H; clear H; intros; unfold open_set; intros; unfold included, interior in H; unfold included in H0; apply (H _ H1). Qed.

forall D : R -> Prop, closed_set D <-> D =_D adherence D

forall D : R -> Prop, closed_set D <-> D =_D adherence D
D:R -> Prop

closed_set D -> D =_D adherence D
D:R -> Prop
D =_D adherence D -> closed_set D
D:R -> Prop
H:closed_set D

included D (adherence D)
D:R -> Prop
H:closed_set D
included (adherence D) D
D:R -> Prop
D =_D adherence D -> closed_set D
D:R -> Prop
H:closed_set D

included (adherence D) D
D:R -> Prop
D =_D adherence D -> closed_set D
D:R -> Prop

D =_D adherence D -> closed_set D
D:R -> Prop
H:(forall x0 : R, D x0 -> adherence D x0) /\ (forall x0 : R, adherence D x0 -> D x0)
H0:forall x0 : R, complementary (adherence D) x0 -> neighbourhood (complementary (adherence D)) x0
x:R
H1:complementary D x

complementary (adherence D) x
D:R -> Prop
H:(forall x0 : R, D x0 -> adherence D x0) /\ (forall x0 : R, adherence D x0 -> D x0)
H0:forall x0 : R, complementary (adherence D) x0 -> neighbourhood (complementary (adherence D)) x0
x:R
H1:complementary D x
H2:complementary (adherence D) x
neighbourhood (complementary D) x
D:R -> Prop
H:(forall x0 : R, D x0 -> adherence D x0) /\ (forall x0 : R, adherence D x0 -> D x0)
H0:forall x0 : R, complementary (adherence D) x0 -> neighbourhood (complementary (adherence D)) x0
x:R
H1:complementary D x
H2:complementary (adherence D) x

neighbourhood (complementary D) x
assert (H3 := H0 _ H2); unfold neighbourhood; unfold neighbourhood in H3; elim H3; intros; exists x0; unfold included; unfold included in H4; intros; assert (H6 := H4 _ H5); unfold complementary in H6; unfold complementary; red; intro; elim H; clear H; intros H _; elim H6; apply (H _ H7). Qed.

forall (D1 D2 : R -> Prop) (x : R), included D1 D2 -> neighbourhood D1 x -> neighbourhood D2 x

forall (D1 D2 : R -> Prop) (x : R), included D1 D2 -> neighbourhood D1 x -> neighbourhood D2 x
unfold included, neighbourhood; intros; elim H0; intros; exists x0; intros; unfold included; unfold included in H1; intros; apply (H _ (H1 _ H2)). Qed.

forall D1 D2 : R -> Prop, open_set D1 -> open_set D2 -> open_set (union_domain D1 D2)

forall D1 D2 : R -> Prop, open_set D1 -> open_set D2 -> open_set (union_domain D1 D2)
D1, D2:R -> Prop
H:forall x0 : R, D1 x0 -> neighbourhood D1 x0
H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0
x:R
H1:D1 x \/ D2 x
H2:D1 x

neighbourhood (union_domain D1 D2) x
D1, D2:R -> Prop
H:forall x0 : R, D1 x0 -> neighbourhood D1 x0
H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0
x:R
H1:D1 x \/ D2 x
H2:D2 x
neighbourhood (union_domain D1 D2) x
D1, D2:R -> Prop
H:forall x0 : R, D1 x0 -> neighbourhood D1 x0
H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0
x:R
H1:D1 x \/ D2 x
H2:D1 x

included D1 (union_domain D1 D2)
D1, D2:R -> Prop
H:forall x0 : R, D1 x0 -> neighbourhood D1 x0
H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0
x:R
H1:D1 x \/ D2 x
H2:D1 x
neighbourhood D1 x
D1, D2:R -> Prop
H:forall x0 : R, D1 x0 -> neighbourhood D1 x0
H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0
x:R
H1:D1 x \/ D2 x
H2:D2 x
neighbourhood (union_domain D1 D2) x
D1, D2:R -> Prop
H:forall x0 : R, D1 x0 -> neighbourhood D1 x0
H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0
x:R
H1:D1 x \/ D2 x
H2:D1 x

neighbourhood D1 x
D1, D2:R -> Prop
H:forall x0 : R, D1 x0 -> neighbourhood D1 x0
H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0
x:R
H1:D1 x \/ D2 x
H2:D2 x
neighbourhood (union_domain D1 D2) x
D1, D2:R -> Prop
H:forall x0 : R, D1 x0 -> neighbourhood D1 x0
H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0
x:R
H1:D1 x \/ D2 x
H2:D2 x

neighbourhood (union_domain D1 D2) x
D1, D2:R -> Prop
H:forall x0 : R, D1 x0 -> neighbourhood D1 x0
H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0
x:R
H1:D1 x \/ D2 x
H2:D2 x

included D2 (union_domain D1 D2)
D1, D2:R -> Prop
H:forall x0 : R, D1 x0 -> neighbourhood D1 x0
H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0
x:R
H1:D1 x \/ D2 x
H2:D2 x
neighbourhood D2 x
D1, D2:R -> Prop
H:forall x0 : R, D1 x0 -> neighbourhood D1 x0
H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0
x:R
H1:D1 x \/ D2 x
H2:D2 x

neighbourhood D2 x
apply H0; assumption. Qed.

forall D1 D2 : R -> Prop, open_set D1 -> open_set D2 -> open_set (intersection_domain D1 D2)

forall D1 D2 : R -> Prop, open_set D1 -> open_set D2 -> open_set (intersection_domain D1 D2)
D1, D2:R -> Prop
H:forall x0 : R, D1 x0 -> neighbourhood D1 x0
H0:forall x0 : R, D2 x0 -> neighbourhood D2 x0
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x

neighbourhood (intersection_domain D1 D2) x
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:included (disc x del1) D1
del2:posreal
H0:included (disc x del2) D2

0 < Rmin del1 del2 -> neighbourhood (fun c : R => D1 c /\ D2 c) x
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:included (disc x del1) D1
del2:posreal
H0:included (disc x del2) D2
0 < Rmin del1 del2
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:included (disc x del1) D1
del2:posreal
H0:included (disc x del2) D2
H6:0 < Rmin del1 del2
del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posreal

neighbourhood (fun c : R => D1 c /\ D2 c) x
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:included (disc x del1) D1
del2:posreal
H0:included (disc x del2) D2
0 < Rmin del1 del2
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1
del2:posreal
H0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1
H6:0 < Rmin del1 del2
del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posreal
x0:R
H7:Rabs (x0 - x) < del

D1 x0 /\ D2 x0
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:included (disc x del1) D1
del2:posreal
H0:included (disc x del2) D2
0 < Rmin del1 del2
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1
del2:posreal
H0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1
H6:0 < Rmin del1 del2
del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posreal
x0:R
H7:Rabs (x0 - x) < del

D1 x0
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1
del2:posreal
H0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1
H6:0 < Rmin del1 del2
del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posreal
x0:R
H7:Rabs (x0 - x) < del
D2 x0
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:included (disc x del1) D1
del2:posreal
H0:included (disc x del2) D2
0 < Rmin del1 del2
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1
del2:posreal
H0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1
H6:0 < Rmin del1 del2
del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posreal
x0:R
H7:Rabs (x0 - x) < del

Rabs (x0 - x) < del
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1
del2:posreal
H0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1
H6:0 < Rmin del1 del2
del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posreal
x0:R
H7:Rabs (x0 - x) < del
del <= del1
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1
del2:posreal
H0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1
H6:0 < Rmin del1 del2
del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posreal
x0:R
H7:Rabs (x0 - x) < del
D2 x0
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:included (disc x del1) D1
del2:posreal
H0:included (disc x del2) D2
0 < Rmin del1 del2
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1
del2:posreal
H0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1
H6:0 < Rmin del1 del2
del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posreal
x0:R
H7:Rabs (x0 - x) < del

del <= del1
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1
del2:posreal
H0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1
H6:0 < Rmin del1 del2
del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posreal
x0:R
H7:Rabs (x0 - x) < del
D2 x0
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:included (disc x del1) D1
del2:posreal
H0:included (disc x del2) D2
0 < Rmin del1 del2
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1
del2:posreal
H0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1
H6:0 < Rmin del1 del2
del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posreal
x0:R
H7:Rabs (x0 - x) < del

D2 x0
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:included (disc x del1) D1
del2:posreal
H0:included (disc x del2) D2
0 < Rmin del1 del2
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1
del2:posreal
H0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1
H6:0 < Rmin del1 del2
del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posreal
x0:R
H7:Rabs (x0 - x) < del

Rabs (x0 - x) < del
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1
del2:posreal
H0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1
H6:0 < Rmin del1 del2
del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posreal
x0:R
H7:Rabs (x0 - x) < del
del <= del2
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:included (disc x del1) D1
del2:posreal
H0:included (disc x del2) D2
0 < Rmin del1 del2
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:forall x1 : R, Rabs (x1 - x) < del1 -> D1 x1
del2:posreal
H0:forall x1 : R, Rabs (x1 - x) < del2 -> D2 x1
H6:0 < Rmin del1 del2
del:={| pos := Rmin del1 del2; cond_pos := H6 |}:posreal
x0:R
H7:Rabs (x0 - x) < del

del <= del2
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:included (disc x del1) D1
del2:posreal
H0:included (disc x del2) D2
0 < Rmin del1 del2
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:included (disc x del1) D1
del2:posreal
H0:included (disc x del2) D2

0 < Rmin del1 del2
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:included (disc x del1) D1
del2:posreal
H0:included (disc x del2) D2
r:del1 <= del2

0 < del1
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:included (disc x del1) D1
del2:posreal
H0:included (disc x del2) D2
n:~ del1 <= del2
0 < del2
D1, D2:R -> Prop
x:R
H1:D1 x /\ D2 x
H2:D1 x
H3:D2 x
H4:exists delta : posreal, included (disc x delta) D1
H5:exists delta : posreal, included (disc x delta) D2
del1:posreal
H:included (disc x del1) D1
del2:posreal
H0:included (disc x del2) D2
n:~ del1 <= del2

0 < del2
apply (cond_pos del2). Qed.

open_set (fun _ : R => False)

open_set (fun _ : R => False)
unfold open_set; intros; elim H. Qed.

open_set (fun _ : R => True)

open_set (fun _ : R => True)
x:R
H:True

exists delta : posreal, included (disc x delta) (fun _ : R => True)
exists (mkposreal 1 Rlt_0_1); unfold included; intros; trivial. Qed.

forall (x : R) (del : posreal), open_set (disc x del)

forall (x : R) (del : posreal), open_set (disc x del)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)

open_set (disc x del)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)

disc x del =_D interior (disc x del)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)

included (disc x del) (interior (disc x del))
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
included (interior (disc x del)) (disc x del)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
x0:R
H2:Rabs (x0 - x) < del

0 < del - Rabs (x - x0) -> neighbourhood (fun y : R => Rabs (y - x) < del) x0
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
x0:R
H2:Rabs (x0 - x) < del
0 < del - Rabs (x - x0)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
included (interior (disc x del)) (disc x del)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
x0:R
H2:Rabs (x0 - x) < del
H3:0 < del - Rabs (x - x0)
del2:={| pos := del - Rabs (x - x0); cond_pos := H3 |}:posreal

neighbourhood (fun y : R => Rabs (y - x) < del) x0
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
x0:R
H2:Rabs (x0 - x) < del
0 < del - Rabs (x - x0)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
included (interior (disc x del)) (disc x del)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
x0:R
H2:Rabs (x0 - x) < del
H3:0 < del - Rabs (x - x0)
del2:={| pos := del - Rabs (x - x0); cond_pos := H3 |}:posreal
x1:R
H4:disc x0 del2 x1

Rabs (x1 - x) < del
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
x0:R
H2:Rabs (x0 - x) < del
0 < del - Rabs (x - x0)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
included (interior (disc x del)) (disc x del)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
x0:R
H2:Rabs (x0 - x) < del
H3:0 < del - Rabs (x - x0)
del2:={| pos := del - Rabs (x - x0); cond_pos := H3 |}:posreal
x1:R
H4:disc x0 del2 x1

Rabs (x1 - x) <= Rabs (x1 - x0) + Rabs (x0 - x)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
x0:R
H2:Rabs (x0 - x) < del
H3:0 < del - Rabs (x - x0)
del2:={| pos := del - Rabs (x - x0); cond_pos := H3 |}:posreal
x1:R
H4:disc x0 del2 x1
Rabs (x1 - x0) + Rabs (x0 - x) < del
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
x0:R
H2:Rabs (x0 - x) < del
0 < del - Rabs (x - x0)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
included (interior (disc x del)) (disc x del)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
x0:R
H2:Rabs (x0 - x) < del
H3:0 < del - Rabs (x - x0)
del2:={| pos := del - Rabs (x - x0); cond_pos := H3 |}:posreal
x1:R
H4:disc x0 del2 x1

Rabs (x1 - x0) + Rabs (x0 - x) < del
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
x0:R
H2:Rabs (x0 - x) < del
0 < del - Rabs (x - x0)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
included (interior (disc x del)) (disc x del)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
x0:R
H2:Rabs (x0 - x) < del
H3:0 < del - Rabs (x - x0)
del2:={| pos := del - Rabs (x - x0); cond_pos := H3 |}:posreal
x1:R
H4:disc x0 del2 x1

Rabs (x1 - x0) + Rabs (x0 - x) < del2 + Rabs (x0 - x)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
x0:R
H2:Rabs (x0 - x) < del
H3:0 < del - Rabs (x - x0)
del2:={| pos := del - Rabs (x - x0); cond_pos := H3 |}:posreal
x1:R
H4:disc x0 del2 x1
del2 + Rabs (x0 - x) = del
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
x0:R
H2:Rabs (x0 - x) < del
0 < del - Rabs (x - x0)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
included (interior (disc x del)) (disc x del)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
x0:R
H2:Rabs (x0 - x) < del
H3:0 < del - Rabs (x - x0)
del2:={| pos := del - Rabs (x - x0); cond_pos := H3 |}:posreal
x1:R
H4:disc x0 del2 x1

Rabs (x1 - x0) < del2
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
x0:R
H2:Rabs (x0 - x) < del
H3:0 < del - Rabs (x - x0)
del2:={| pos := del - Rabs (x - x0); cond_pos := H3 |}:posreal
x1:R
H4:disc x0 del2 x1
del2 + Rabs (x0 - x) = del
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
x0:R
H2:Rabs (x0 - x) < del
0 < del - Rabs (x - x0)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
included (interior (disc x del)) (disc x del)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
x0:R
H2:Rabs (x0 - x) < del
H3:0 < del - Rabs (x - x0)
del2:={| pos := del - Rabs (x - x0); cond_pos := H3 |}:posreal
x1:R
H4:disc x0 del2 x1

del2 + Rabs (x0 - x) = del
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
x0:R
H2:Rabs (x0 - x) < del
0 < del - Rabs (x - x0)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
included (interior (disc x del)) (disc x del)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
x0:R
H2:Rabs (x0 - x) < del

0 < del - Rabs (x - x0)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)
included (interior (disc x del)) (disc x del)
x:R
del:posreal
H:open_set (disc x del) <-> disc x del =_D interior (disc x del)
H0:open_set (disc x del) -> disc x del =_D interior (disc x del)
H1:disc x del =_D interior (disc x del) -> open_set (disc x del)

included (interior (disc x del)) (disc x del)
apply interior_P1. Qed.

forall (f : R -> R) (x : R), continuity_pt f x <-> (forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y)))

forall (f : R -> R) (x : R), continuity_pt f x <-> (forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y)))
f:R -> R
x:R

continuity_pt f x -> forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))
f:R -> R
x:R
(forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))) -> continuity_pt f x
f:R -> R
x:R
H:continuity_pt f x
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W

exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))
f:R -> R
x:R
(forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))) -> continuity_pt f x
f:R -> R
x:R
H:continuity_pt f x
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:included (disc (f x) del1) W

exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))
f:R -> R
x:R
(forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))) -> continuity_pt f x
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:included (disc (f x) del1) W

exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))
f:R -> R
x:R
(forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))) -> continuity_pt f x
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:included (disc (f x) del1) W
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)

exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))
f:R -> R
x:R
(forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))) -> continuity_pt f x
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:included (disc (f x) del1) W
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)
del2:R
H3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)

exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))
f:R -> R
x:R
(forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))) -> continuity_pt f x
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:included (disc (f x) del1) W
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)
del2:R
H3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)
H4:del2 > 0
H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1

exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))
f:R -> R
x:R
(forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))) -> continuity_pt f x
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:included (disc (f x) del1) W
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)
del2:R
H3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)
H4:del2 > 0
H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1

neighbourhood (disc x {| pos := del2; cond_pos := H4 |}) x /\ (forall y : R, disc x {| pos := del2; cond_pos := H4 |} y -> W (f y))
f:R -> R
x:R
(forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))) -> continuity_pt f x
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:forall x0 : R, disc (f x) del1 x0 -> W x0
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)
del2:R
H3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)
H4:del2 > 0
H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1

neighbourhood (disc x {| pos := del2; cond_pos := H4 |}) x
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:forall x0 : R, disc (f x) del1 x0 -> W x0
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)
del2:R
H3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)
H4:del2 > 0
H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1
forall y : R, disc x {| pos := del2; cond_pos := H4 |} y -> W (f y)
f:R -> R
x:R
(forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))) -> continuity_pt f x
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:forall x0 : R, disc (f x) del1 x0 -> W x0
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)
del2:R
H3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)
H4:del2 > 0
H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1

exists delta : posreal, included (fun y : R => Rabs (y - x) < delta) (fun y : R => Rabs (y - x) < {| pos := del2; cond_pos := H4 |})
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:forall x0 : R, disc (f x) del1 x0 -> W x0
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)
del2:R
H3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)
H4:del2 > 0
H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1
forall y : R, disc x {| pos := del2; cond_pos := H4 |} y -> W (f y)
f:R -> R
x:R
(forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))) -> continuity_pt f x
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:forall x0 : R, disc (f x) del1 x0 -> W x0
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)
del2:R
H3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)
H4:del2 > 0
H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1

included (fun y : R => Rabs (y - x) < {| pos := del2; cond_pos := H4 |}) (fun y : R => Rabs (y - x) < {| pos := del2; cond_pos := H4 |})
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:forall x0 : R, disc (f x) del1 x0 -> W x0
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)
del2:R
H3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)
H4:del2 > 0
H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1
forall y : R, disc x {| pos := del2; cond_pos := H4 |} y -> W (f y)
f:R -> R
x:R
(forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))) -> continuity_pt f x
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:forall x0 : R, disc (f x) del1 x0 -> W x0
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)
del2:R
H3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)
H4:del2 > 0
H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1

forall y : R, disc x {| pos := del2; cond_pos := H4 |} y -> W (f y)
f:R -> R
x:R
(forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))) -> continuity_pt f x
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:forall x0 : R, disc (f x) del1 x0 -> W x0
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)
del2:R
H3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)
H4:del2 > 0
H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1
y:R
H6:disc x {| pos := del2; cond_pos := H4 |} y
H7:y = x

Rabs (f y - f x) < del1
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:forall x0 : R, disc (f x) del1 x0 -> W x0
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)
del2:R
H3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)
H4:del2 > 0
H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1
y:R
H6:disc x {| pos := del2; cond_pos := H4 |} y
H7:y <> x
Rabs (f y - f x) < del1
f:R -> R
x:R
(forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))) -> continuity_pt f x
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:forall x0 : R, disc (f x) del1 x0 -> W x0
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)
del2:R
H3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)
H4:del2 > 0
H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1
y:R
H6:disc x {| pos := del2; cond_pos := H4 |} y
H7:y <> x

Rabs (f y - f x) < del1
f:R -> R
x:R
(forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))) -> continuity_pt f x
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:forall x0 : R, disc (f x) del1 x0 -> W x0
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)
del2:R
H3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)
H4:del2 > 0
H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1
y:R
H6:disc x {| pos := del2; cond_pos := H4 |} y
H7:y <> x

D_x no_cond x y
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:forall x0 : R, disc (f x) del1 x0 -> W x0
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)
del2:R
H3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)
H4:del2 > 0
H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1
y:R
H6:disc x {| pos := del2; cond_pos := H4 |} y
H7:y <> x
Rabs (y - x) < del2
f:R -> R
x:R
(forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))) -> continuity_pt f x
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:forall x0 : R, disc (f x) del1 x0 -> W x0
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)
del2:R
H3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)
H4:del2 > 0
H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1
y:R
H6:disc x {| pos := del2; cond_pos := H4 |} y
H7:y <> x

True
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:forall x0 : R, disc (f x) del1 x0 -> W x0
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)
del2:R
H3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)
H4:del2 > 0
H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1
y:R
H6:disc x {| pos := del2; cond_pos := H4 |} y
H7:y <> x
x <> y
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:forall x0 : R, disc (f x) del1 x0 -> W x0
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)
del2:R
H3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)
H4:del2 > 0
H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1
y:R
H6:disc x {| pos := del2; cond_pos := H4 |} y
H7:y <> x
Rabs (y - x) < del2
f:R -> R
x:R
(forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))) -> continuity_pt f x
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:forall x0 : R, disc (f x) del1 x0 -> W x0
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)
del2:R
H3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)
H4:del2 > 0
H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1
y:R
H6:disc x {| pos := del2; cond_pos := H4 |} y
H7:y <> x

x <> y
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:forall x0 : R, disc (f x) del1 x0 -> W x0
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)
del2:R
H3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)
H4:del2 > 0
H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1
y:R
H6:disc x {| pos := del2; cond_pos := H4 |} y
H7:y <> x
Rabs (y - x) < del2
f:R -> R
x:R
(forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))) -> continuity_pt f x
f:R -> R
x:R
H:forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
W:R -> Prop
H0:exists delta : posreal, included (disc (f x) delta) W
del1:posreal
H1:forall x0 : R, disc (f x) del1 x0 -> W x0
H2:exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < del1)
del2:R
H3:del2 > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1)
H4:del2 > 0
H5:forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del2 -> Rabs (f x0 - f x) < del1
y:R
H6:disc x {| pos := del2; cond_pos := H4 |} y
H7:y <> x

Rabs (y - x) < del2
f:R -> R
x:R
(forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))) -> continuity_pt f x
f:R -> R
x:R

(forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))) -> continuity_pt f x
f:R -> R
x:R
H:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))
eps:R
H0:eps > 0

exists alp : R, alp > 0 /\ (forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < alp -> dist R_met (f x0) (f x) < eps)
f:R -> R
x:R
H:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))
eps:R
H0:eps > 0
H1:neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))

exists alp : R, alp > 0 /\ (forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < alp -> dist R_met (f x0) (f x) < eps)
f:R -> R
x:R
H:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))
eps:R
H0:eps > 0
H1:neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))

neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x) -> exists alp : R, alp > 0 /\ (forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < alp -> dist R_met (f x0) (f x) < eps)
f:R -> R
x:R
H:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))
eps:R
H0:eps > 0
H1:neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))
neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x)
f:R -> R
x:R
H:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))
eps:R
H0:eps > 0
H1:neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))
H2:neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x)
H3:exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))

exists alp : R, alp > 0 /\ (forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < alp -> dist R_met (f x0) (f x) < eps)
f:R -> R
x:R
H:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))
eps:R
H0:eps > 0
H1:neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))
neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x)
f:R -> R
x:R
H:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))
eps:R
H0:eps > 0
H1:neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))
H2:neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x)
H3:exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))
D:R -> Prop
H4:neighbourhood D x /\ (forall y : R, D y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))
H5:exists delta : posreal, included (disc x delta) D
H6:forall y : R, D y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y)
del1:posreal
H7:included (disc x del1) D

exists alp : R, alp > 0 /\ (forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < alp -> dist R_met (f x0) (f x) < eps)
f:R -> R
x:R
H:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))
eps:R
H0:eps > 0
H1:neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))
neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x)
f:R -> R
x:R
H:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))
eps:R
H0:eps > 0
H1:neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))
H2:neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x)
H3:exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))
D:R -> Prop
H4:neighbourhood D x /\ (forall y : R, D y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))
H5:exists delta : posreal, included (disc x delta) D
H6:forall y : R, D y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y)
del1:posreal
H7:included (disc x del1) D

del1 > 0
f:R -> R
x:R
H:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))
eps:R
H0:eps > 0
H1:neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))
H2:neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x)
H3:exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))
D:R -> Prop
H4:neighbourhood D x /\ (forall y : R, D y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))
H5:exists delta : posreal, included (disc x delta) D
H6:forall y : R, D y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y)
del1:posreal
H7:included (disc x del1) D
forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < del1 -> dist R_met (f x0) (f x) < eps
f:R -> R
x:R
H:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))
eps:R
H0:eps > 0
H1:neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))
neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x)
f:R -> R
x:R
H:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))
eps:R
H0:eps > 0
H1:neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))
H2:neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x)
H3:exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))
D:R -> Prop
H4:neighbourhood D x /\ (forall y : R, D y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))
H5:exists delta : posreal, included (disc x delta) D
H6:forall y : R, D y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y)
del1:posreal
H7:included (disc x del1) D

forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < del1 -> dist R_met (f x0) (f x) < eps
f:R -> R
x:R
H:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))
eps:R
H0:eps > 0
H1:neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))
neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x)
f:R -> R
x:R
H:forall W : R -> Prop, neighbourhood W (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> W (f y))
eps:R
H0:eps > 0
H1:neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x) -> exists V : R -> Prop, neighbourhood V x /\ (forall y : R, V y -> disc (f x) {| pos := eps; cond_pos := H0 |} (f y))

neighbourhood (disc (f x) {| pos := eps; cond_pos := H0 |}) (f x)
unfold neighbourhood, disc; exists (mkposreal eps H0); unfold included; intros; assumption. Qed. Definition image_rec (f:R -> R) (D:R -> Prop) (x:R) : Prop := D (f x). (**********)

forall (f : R -> R) (D : R -> Prop), continuity f -> open_set D -> open_set (image_rec f D)

forall (f : R -> R) (D : R -> Prop), continuity f -> open_set D -> open_set (image_rec f D)
intros; unfold open_set in H0; unfold open_set; intros; assert (H2 := continuity_P1 f x); elim H2; intros H3 _; assert (H4 := H3 (H x)); unfold neighbourhood, image_rec; unfold image_rec in H1; assert (H5 := H4 D (H0 (f x) H1)); elim H5; intros V0 H6; elim H6; intros; unfold neighbourhood in H7; elim H7; intros del H9; exists del; unfold included in H9; unfold included; intros; apply (H8 _ (H9 _ H10)). Qed. (**********)

forall f : R -> R, continuity f <-> (forall D : R -> Prop, open_set D -> open_set (image_rec f D))

forall f : R -> R, continuity f <-> (forall D : R -> Prop, open_set D -> open_set (image_rec f D))
f:R -> R

continuity f -> forall D : R -> Prop, open_set D -> open_set (image_rec f D)
f:R -> R
(forall D : R -> Prop, open_set D -> open_set (image_rec f D)) -> continuity f
f:R -> R

(forall D : R -> Prop, open_set D -> open_set (image_rec f D)) -> continuity f
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0

open_set (disc (f x) {| pos := eps; cond_pos := H0 |}) -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0
open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0
H1:open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
H2:open_set (image_rec f (disc (f x) {| pos := eps; cond_pos := H0 |}))

exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0
open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0
H1:open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
H2:forall x0 : R, disc (f x) {| pos := eps; cond_pos := H0 |} (f x0) -> neighbourhood (fun x1 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x1)) x0

disc (f x) {| pos := eps; cond_pos := H0 |} (f x) -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0
H1:open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
H2:forall x0 : R, disc (f x) {| pos := eps; cond_pos := H0 |} (f x0) -> neighbourhood (fun x1 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x1)) x0
disc (f x) {| pos := eps; cond_pos := H0 |} (f x)
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0
open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0
H1:open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
H2:forall x0 : R, disc (f x) {| pos := eps; cond_pos := H0 |} (f x0) -> neighbourhood (fun x1 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x1)) x0
H3:disc (f x) {| pos := eps; cond_pos := H0 |} (f x)
H4:neighbourhood (fun x0 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x0)) x

exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0
H1:open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
H2:forall x0 : R, disc (f x) {| pos := eps; cond_pos := H0 |} (f x0) -> neighbourhood (fun x1 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x1)) x0
disc (f x) {| pos := eps; cond_pos := H0 |} (f x)
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0
open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0
H1:open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
H2:forall x0 : R, disc (f x) {| pos := eps; cond_pos := H0 |} (f x0) -> neighbourhood (fun x1 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x1)) x0
H3:disc (f x) {| pos := eps; cond_pos := H0 |} (f x)
H4:exists delta : posreal, included (disc x delta) (fun x0 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x0))
del:posreal
H5:included (disc x del) (fun x0 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x0))

exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < alp -> Rabs (f x0 - f x) < eps)
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0
H1:open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
H2:forall x0 : R, disc (f x) {| pos := eps; cond_pos := H0 |} (f x0) -> neighbourhood (fun x1 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x1)) x0
disc (f x) {| pos := eps; cond_pos := H0 |} (f x)
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0
open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0
H1:open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
H2:forall x0 : R, disc (f x) {| pos := eps; cond_pos := H0 |} (f x0) -> neighbourhood (fun x1 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x1)) x0
H3:disc (f x) {| pos := eps; cond_pos := H0 |} (f x)
H4:exists delta : posreal, included (disc x delta) (fun x0 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x0))
del:posreal
H5:included (disc x del) (fun x0 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x0))

del > 0
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0
H1:open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
H2:forall x0 : R, disc (f x) {| pos := eps; cond_pos := H0 |} (f x0) -> neighbourhood (fun x1 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x1)) x0
H3:disc (f x) {| pos := eps; cond_pos := H0 |} (f x)
H4:exists delta : posreal, included (disc x delta) (fun x0 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x0))
del:posreal
H5:included (disc x del) (fun x0 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x0))
forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del -> Rabs (f x0 - f x) < eps
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0
H1:open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
H2:forall x0 : R, disc (f x) {| pos := eps; cond_pos := H0 |} (f x0) -> neighbourhood (fun x1 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x1)) x0
disc (f x) {| pos := eps; cond_pos := H0 |} (f x)
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0
open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0
H1:open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
H2:forall x0 : R, disc (f x) {| pos := eps; cond_pos := H0 |} (f x0) -> neighbourhood (fun x1 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x1)) x0
H3:disc (f x) {| pos := eps; cond_pos := H0 |} (f x)
H4:exists delta : posreal, included (disc x delta) (fun x0 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x0))
del:posreal
H5:included (disc x del) (fun x0 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x0))

forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < del -> Rabs (f x0 - f x) < eps
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0
H1:open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
H2:forall x0 : R, disc (f x) {| pos := eps; cond_pos := H0 |} (f x0) -> neighbourhood (fun x1 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x1)) x0
disc (f x) {| pos := eps; cond_pos := H0 |} (f x)
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0
open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0
H1:open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
H2:forall x0 : R, disc (f x) {| pos := eps; cond_pos := H0 |} (f x0) -> neighbourhood (fun x1 : R => disc (f x) {| pos := eps; cond_pos := H0 |} (f x1)) x0

disc (f x) {| pos := eps; cond_pos := H0 |} (f x)
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0
open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
f:R -> R
H:forall D : R -> Prop, open_set D -> open_set (image_rec f D)
x, eps:R
H0:eps > 0

open_set (disc (f x) {| pos := eps; cond_pos := H0 |})
apply disc_P1. Qed. (**********)

forall x y : R, x <> y -> exists V W : R -> Prop, neighbourhood V x /\ neighbourhood W y /\ ~ (exists y0 : R, intersection_domain V W y0)

forall x y : R, x <> y -> exists V W : R -> Prop, neighbourhood V x /\ neighbourhood W y /\ ~ (exists y0 : R, intersection_domain V W y0)
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R

exists V W : R -> Prop, neighbourhood V x /\ neighbourhood W y /\ ~ (exists y0 : R, intersection_domain V W y0)
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R

0 < D / 2 -> exists V W : R -> Prop, neighbourhood V x /\ neighbourhood W y /\ ~ (exists y0 : R, intersection_domain V W y0)
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
0 < D / 2
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
H:0 < D / 2

exists W : R -> Prop, neighbourhood (disc x {| pos := D / 2; cond_pos := H |}) x /\ neighbourhood W y /\ ~ (exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) W y0)
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
0 < D / 2
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
H:0 < D / 2

neighbourhood (disc x {| pos := D / 2; cond_pos := H |}) x
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
H:0 < D / 2
neighbourhood (disc y {| pos := D / 2; cond_pos := H |}) y /\ ~ (exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0)
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
0 < D / 2
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
H:0 < D / 2

neighbourhood (disc y {| pos := D / 2; cond_pos := H |}) y /\ ~ (exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0)
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
0 < D / 2
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
H:0 < D / 2

neighbourhood (disc y {| pos := D / 2; cond_pos := H |}) y
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
H:0 < D / 2
~ (exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0)
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
0 < D / 2
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
H:0 < D / 2

~ (exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0)
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
0 < D / 2
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
H:0 < D / 2
H0:exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0
x0:R
H1:disc x {| pos := D / 2; cond_pos := H |} x0 /\ disc y {| pos := D / 2; cond_pos := H |} x0
H2:disc x {| pos := D / 2; cond_pos := H |} x0
H3:disc y {| pos := D / 2; cond_pos := H |} x0

False
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
0 < D / 2
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
H:0 < D / 2
H0:exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0
x0:R
H1:disc x {| pos := D / 2; cond_pos := H |} x0 /\ disc y {| pos := D / 2; cond_pos := H |} x0
H2:disc x {| pos := D / 2; cond_pos := H |} x0
H3:disc y {| pos := D / 2; cond_pos := H |} x0

D < D -> False
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
H:0 < D / 2
H0:exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0
x0:R
H1:disc x {| pos := D / 2; cond_pos := H |} x0 /\ disc y {| pos := D / 2; cond_pos := H |} x0
H2:disc x {| pos := D / 2; cond_pos := H |} x0
H3:disc y {| pos := D / 2; cond_pos := H |} x0
D < D
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
0 < D / 2
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
H:0 < D / 2
H0:exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0
x0:R
H1:disc x {| pos := D / 2; cond_pos := H |} x0 /\ disc y {| pos := D / 2; cond_pos := H |} x0
H2:disc x {| pos := D / 2; cond_pos := H |} x0
H3:disc y {| pos := D / 2; cond_pos := H |} x0

D < D
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
0 < D / 2
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
H:0 < D / 2
H0:exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0
x0:R
H1:disc x {| pos := D / 2; cond_pos := H |} x0 /\ disc y {| pos := D / 2; cond_pos := H |} x0
H2:disc x {| pos := D / 2; cond_pos := H |} x0
H3:disc y {| pos := D / 2; cond_pos := H |} x0

Rabs (x - y) <= Rabs (x - x0) + Rabs (x0 - y)
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
H:0 < D / 2
H0:exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0
x0:R
H1:disc x {| pos := D / 2; cond_pos := H |} x0 /\ disc y {| pos := D / 2; cond_pos := H |} x0
H2:disc x {| pos := D / 2; cond_pos := H |} x0
H3:disc y {| pos := D / 2; cond_pos := H |} x0
Rabs (x - x0) + Rabs (x0 - y) < D
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
0 < D / 2
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
H:0 < D / 2
H0:exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0
x0:R
H1:disc x {| pos := D / 2; cond_pos := H |} x0 /\ disc y {| pos := D / 2; cond_pos := H |} x0
H2:disc x {| pos := D / 2; cond_pos := H |} x0
H3:disc y {| pos := D / 2; cond_pos := H |} x0

Rabs (x - x0) + Rabs (x0 - y) < D
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
0 < D / 2
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
H:0 < D / 2
H0:exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0
x0:R
H1:disc x {| pos := D / 2; cond_pos := H |} x0 /\ disc y {| pos := D / 2; cond_pos := H |} x0
H2:disc x {| pos := D / 2; cond_pos := H |} x0
H3:disc y {| pos := D / 2; cond_pos := H |} x0

Rabs (x - x0) < D / 2
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
H:0 < D / 2
H0:exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0
x0:R
H1:disc x {| pos := D / 2; cond_pos := H |} x0 /\ disc y {| pos := D / 2; cond_pos := H |} x0
H2:disc x {| pos := D / 2; cond_pos := H |} x0
H3:disc y {| pos := D / 2; cond_pos := H |} x0
Rabs (x0 - y) < D / 2
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
0 < D / 2
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
H:0 < D / 2
H0:exists y0 : R, intersection_domain (disc x {| pos := D / 2; cond_pos := H |}) (disc y {| pos := D / 2; cond_pos := H |}) y0
x0:R
H1:disc x {| pos := D / 2; cond_pos := H |} x0 /\ disc y {| pos := D / 2; cond_pos := H |} x0
H2:disc x {| pos := D / 2; cond_pos := H |} x0
H3:disc y {| pos := D / 2; cond_pos := H |} x0

Rabs (x0 - y) < D / 2
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
0 < D / 2
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R

0 < D / 2
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R

0 < D
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R
0 < / 2
x, y:R
Hsep:x <> y
D:=Rabs (x - y):R

0 < / 2
apply Rinv_0_lt_compat; prove_sup0. Qed. Record family : Type := mkfamily {ind : R -> Prop; f :> R -> R -> Prop; cond_fam : forall x:R, (exists y : R, f x y) -> ind x}. Definition family_open_set (f:family) : Prop := forall x:R, open_set (f x). Definition domain_finite (D:R -> Prop) : Prop := exists l : Rlist, (forall x:R, D x <-> In x l). Definition family_finite (f:family) : Prop := domain_finite (ind f). Definition covering (D:R -> Prop) (f:family) : Prop := forall x:R, D x -> exists y : R, f y x. Definition covering_open_set (D:R -> Prop) (f:family) : Prop := covering D f /\ family_open_set f. Definition covering_finite (D:R -> Prop) (f:family) : Prop := covering D f /\ family_finite f.

forall (f0 : family) (D : R -> Prop) (x : R), (exists y : R, (fun z1 z2 : R => f0 z1 z2 /\ D z1) x y) -> intersection_domain (ind f0) D x

forall (f0 : family) (D : R -> Prop) (x : R), (exists y : R, (fun z1 z2 : R => f0 z1 z2 /\ D z1) x y) -> intersection_domain (ind f0) D x
f0:family
D:R -> Prop
x:R
H:exists y : R, (fun z1 z2 : R => f0 z1 z2 /\ D z1) x y
x0:R
H0:f0 x x0 /\ D x
H1:f0 x x0
H2:D x

ind f0 x
f0:family
D:R -> Prop
x:R
H:exists y : R, (fun z1 z2 : R => f0 z1 z2 /\ D z1) x y
x0:R
H0:f0 x x0 /\ D x
H1:f0 x x0
H2:D x
D x
f0:family
D:R -> Prop
x:R
H:exists y : R, (fun z1 z2 : R => f0 z1 z2 /\ D z1) x y
x0:R
H0:f0 x x0 /\ D x
H1:f0 x x0
H2:D x

D x
assumption. Qed. Definition subfamily (f:family) (D:R -> Prop) : family := mkfamily (intersection_domain (ind f) D) (fun x y:R => f x y /\ D x) (restriction_family f D). Definition compact (X:R -> Prop) : Prop := forall f:family, covering_open_set X f -> exists D : R -> Prop, covering_finite X (subfamily f D). (**********)

forall (f0 : family) (D : R -> Prop), family_open_set f0 -> family_open_set (subfamily f0 D)

forall (f0 : family) (D : R -> Prop), family_open_set f0 -> family_open_set (subfamily f0 D)
f0:family
D:R -> Prop
H:forall x0 : R, open_set (f0 x0)
x:R
H0:D x \/ ~ D x

open_set (fun y : R => f0 x y /\ D x)
f0:family
D:R -> Prop
H:forall x0 : R, open_set (f0 x0)
x:R
H0:D x \/ ~ D x
H1:D x

open_set (fun y : R => f0 x y /\ D x)
f0:family
D:R -> Prop
H:forall x0 : R, open_set (f0 x0)
x:R
H0:D x \/ ~ D x
H1:~ D x
open_set (fun y : R => f0 x y /\ D x)
f0:family
D:R -> Prop
H:forall x0 : R, open_set (f0 x0)
x:R
H0:D x \/ ~ D x
H1:D x

(open_set (f0 x) -> open_set (fun y : R => f0 x y /\ D x)) -> open_set (fun y : R => f0 x y /\ D x)
f0:family
D:R -> Prop
H:forall x0 : R, open_set (f0 x0)
x:R
H0:D x \/ ~ D x
H1:D x
open_set (f0 x) -> open_set (fun y : R => f0 x y /\ D x)
f0:family
D:R -> Prop
H:forall x0 : R, open_set (f0 x0)
x:R
H0:D x \/ ~ D x
H1:~ D x
open_set (fun y : R => f0 x y /\ D x)
f0:family
D:R -> Prop
H:forall x0 : R, open_set (f0 x0)
x:R
H0:D x \/ ~ D x
H1:D x

open_set (f0 x) -> open_set (fun y : R => f0 x y /\ D x)
f0:family
D:R -> Prop
H:forall x0 : R, open_set (f0 x0)
x:R
H0:D x \/ ~ D x
H1:~ D x
open_set (fun y : R => f0 x y /\ D x)
f0:family
D:R -> Prop
H:forall x3 : R, open_set (f0 x3)
x:R
H0:D x \/ ~ D x
H1:D x
H2:forall x3 : R, f0 x x3 -> exists delta : posreal, included (disc x3 delta) (f0 x)
x0:R
H3:f0 x x0 /\ D x
H4:f0 x x0
H5:D x
H6:exists delta : posreal, included (disc x0 delta) (f0 x)
x1:posreal
H7:included (disc x0 x1) (f0 x)
x2:R
H8:disc x0 x1 x2

f0 x x2
f0:family
D:R -> Prop
H:forall x3 : R, open_set (f0 x3)
x:R
H0:D x \/ ~ D x
H1:D x
H2:forall x3 : R, f0 x x3 -> exists delta : posreal, included (disc x3 delta) (f0 x)
x0:R
H3:f0 x x0 /\ D x
H4:f0 x x0
H5:D x
H6:exists delta : posreal, included (disc x0 delta) (f0 x)
x1:posreal
H7:included (disc x0 x1) (f0 x)
x2:R
H8:disc x0 x1 x2
D x
f0:family
D:R -> Prop
H:forall x0 : R, open_set (f0 x0)
x:R
H0:D x \/ ~ D x
H1:~ D x
open_set (fun y : R => f0 x y /\ D x)
f0:family
D:R -> Prop
H:forall x3 : R, open_set (f0 x3)
x:R
H0:D x \/ ~ D x
H1:D x
H2:forall x3 : R, f0 x x3 -> exists delta : posreal, included (disc x3 delta) (f0 x)
x0:R
H3:f0 x x0 /\ D x
H4:f0 x x0
H5:D x
H6:exists delta : posreal, included (disc x0 delta) (f0 x)
x1:posreal
H7:included (disc x0 x1) (f0 x)
x2:R
H8:disc x0 x1 x2

D x
f0:family
D:R -> Prop
H:forall x0 : R, open_set (f0 x0)
x:R
H0:D x \/ ~ D x
H1:~ D x
open_set (fun y : R => f0 x y /\ D x)
f0:family
D:R -> Prop
H:forall x0 : R, open_set (f0 x0)
x:R
H0:D x \/ ~ D x
H1:~ D x

open_set (fun y : R => f0 x y /\ D x)
f0:family
D:R -> Prop
H:forall x0 : R, open_set (f0 x0)
x:R
H0:D x \/ ~ D x
H1:~ D x

(open_set (fun _ : R => False) -> open_set (fun y : R => f0 x y /\ D x)) -> open_set (fun y : R => f0 x y /\ D x)
f0:family
D:R -> Prop
H:forall x0 : R, open_set (f0 x0)
x:R
H0:D x \/ ~ D x
H1:~ D x
open_set (fun _ : R => False) -> open_set (fun y : R => f0 x y /\ D x)
f0:family
D:R -> Prop
H:forall x0 : R, open_set (f0 x0)
x:R
H0:D x \/ ~ D x
H1:~ D x

open_set (fun _ : R => False) -> open_set (fun y : R => f0 x y /\ D x)
unfold open_set; unfold neighbourhood; intros; elim H3; intros; elim H1; assumption. Qed. Definition bounded (D:R -> Prop) : Prop := exists m : R, (exists M : R, (forall x:R, D x -> m <= x <= M)).

forall D1 D2 : R -> Prop, open_set D1 -> D1 =_D D2 -> open_set D2

forall D1 D2 : R -> Prop, open_set D1 -> D1 =_D D2 -> open_set D2
D1, D2:R -> Prop
H:forall x0 : R, D1 x0 -> exists delta : posreal, included (disc x0 delta) D1
H0:D1 =_D D2
x:R
H1:D2 x

exists delta : posreal, included (disc x delta) D2
D1, D2:R -> Prop
H:forall x0 : R, D1 x0 -> exists delta : posreal, included (disc x0 delta) D1
H0:included D1 D2 /\ included D2 D1
x:R
H1:D2 x
H2:included D1 D2
H3:included D2 D1

exists delta : posreal, included (disc x delta) D2
D1, D2:R -> Prop
H:forall x0 : R, D1 x0 -> exists delta : posreal, included (disc x0 delta) D1
H0:included D1 D2 /\ included D2 D1
x:R
H1:D2 x
H2:included D1 D2
H3:included D2 D1
H4:exists delta : posreal, included (disc x delta) D1

exists delta : posreal, included (disc x delta) D2
D1, D2:R -> Prop
H:forall x1 : R, D1 x1 -> exists delta : posreal, included (disc x1 delta) D1
H0:included D1 D2 /\ included D2 D1
x:R
H1:D2 x
H2:included D1 D2
H3:included D2 D1
H4:exists delta : posreal, included (disc x delta) D1
x0:posreal
H5:included (disc x x0) D1

exists delta : posreal, included (disc x delta) D2
exists x0; apply included_trans with D1; assumption. Qed. (**********)

forall X : R -> Prop, compact X -> bounded X

forall X : R -> Prop, compact X -> bounded X
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True

bounded X
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)

covering_open_set X f0 -> bounded X
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:covering X (subfamily f0 D')
H6:exists l0 : Rlist, forall x : R, ind (subfamily f0 D') x <-> In x l0
l:Rlist
H7:forall x : R, ind (subfamily f0 D') x <-> In x l
r:=MaxRlist l:R

exists m M : R, forall x : R, X x -> m <= x <= M
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x0 y : R => Rabs y < x0:R -> R -> Prop
H0:forall x0 : R, (exists y : R, g x0 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:covering X (subfamily f0 D')
H6:exists l0 : Rlist, forall x0 : R, ind (subfamily f0 D') x0 <-> In x0 l0
l:Rlist
H7:forall x0 : R, ind (subfamily f0 D') x0 <-> In x0 l
r:=MaxRlist l:R
x:R
H8:X x

- r <= x <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:g x0 x
H12:D' x0
H13:intersection_domain D D' x0 <-> In x0 l

intersection_domain D D' x0 -> - r <= x <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:g x0 x
H12:D' x0
H13:intersection_domain D D' x0 <-> In x0 l
intersection_domain D D' x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:g x0 x
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0

- r <= x <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:g x0 x
H12:D' x0
H13:intersection_domain D D' x0 <-> In x0 l
intersection_domain D D' x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l

- r <= x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
x <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:g x0 x
H12:D' x0
H13:intersection_domain D D' x0 <-> In x0 l
intersection_domain D D' x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l

x0 <= r -> - r <= x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
x0 <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
x <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:g x0 x
H12:D' x0
H13:intersection_domain D D' x0 <-> In x0 l
intersection_domain D D' x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
H17:x0 <= r

Rabs x < r -> - r <= x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
H17:x0 <= r
Rabs x < r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
x0 <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
x <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:g x0 x
H12:D' x0
H13:intersection_domain D D' x0 <-> In x0 l
intersection_domain D D' x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
H17:x0 <= r

Rabs x < r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
x0 <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
x <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:g x0 x
H12:D' x0
H13:intersection_domain D D' x0 <-> In x0 l
intersection_domain D D' x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l

x0 <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
x <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:g x0 x
H12:D' x0
H13:intersection_domain D D' x0 <-> In x0 l
intersection_domain D D' x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l

x <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:g x0 x
H12:D' x0
H13:intersection_domain D D' x0 <-> In x0 l
intersection_domain D D' x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l

x0 <= r -> x <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
x0 <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:g x0 x
H12:D' x0
H13:intersection_domain D D' x0 <-> In x0 l
intersection_domain D D' x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
H17:x0 <= r

x <= Rabs x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
H17:x0 <= r
Rabs x <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
x0 <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:g x0 x
H12:D' x0
H13:intersection_domain D D' x0 <-> In x0 l
intersection_domain D D' x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
H17:x0 <= r

Rabs x <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
x0 <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:g x0 x
H12:D' x0
H13:intersection_domain D D' x0 <-> In x0 l
intersection_domain D D' x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
H17:x0 <= r

Rabs x <= x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
H17:x0 <= r
x0 <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
x0 <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:g x0 x
H12:D' x0
H13:intersection_domain D D' x0 <-> In x0 l
intersection_domain D D' x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
H17:x0 <= r

x0 <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l
x0 <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:g x0 x
H12:D' x0
H13:intersection_domain D D' x0 <-> In x0 l
intersection_domain D D' x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:Rabs x < x0
H12:D' x0
H13:intersection_domain D D' x0 -> In x0 l
H14:In x0 l -> intersection_domain D D' x0
H15:intersection_domain D D' x0
H16:In x0 l

x0 <= r
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:g x0 x
H12:D' x0
H13:intersection_domain D D' x0 <-> In x0 l
intersection_domain D D' x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x1 y : R => Rabs y < x1:R -> R -> Prop
H0:forall x1 : R, (exists y : R, g x1 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H2:covering_open_set X f0
H3:exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
D':R -> Prop
H4:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H5:forall x1 : R, X x1 -> exists y : R, subfamily f0 D' y x1
H6:exists l0 : Rlist, forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l0
l:Rlist
H7:forall x1 : R, ind (subfamily f0 D') x1 <-> In x1 l
r:=MaxRlist l:R
x:R
H8:X x
H9:exists y : R, subfamily f0 D' y x
x0:R
H10:g x0 x /\ D' x0
H11:g x0 x
H12:D' x0
H13:intersection_domain D D' x0 <-> In x0 l

intersection_domain D D' x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)

covering_open_set X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)

covering X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
family_open_set f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x y : R => Rabs y < x:R -> R -> Prop
H0:forall x : R, (exists y : R, g x y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)

family_open_set f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x0 y : R => Rabs y < x0:R -> R -> Prop
H0:forall x0 : R, (exists y : R, g x0 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H2:0 < x

open_set (f0 x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x0 y : R => Rabs y < x0:R -> R -> Prop
H0:forall x0 : R, (exists y : R, g x0 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H2:0 = x \/ 0 > x
open_set (f0 x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x0 y : R => Rabs y < x0:R -> R -> Prop
H0:forall x0 : R, (exists y : R, g x0 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H2:0 < x

open_set (disc 0 {| pos := x; cond_pos := H2 |})
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x0 y : R => Rabs y < x0:R -> R -> Prop
H0:forall x0 : R, (exists y : R, g x0 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H2:0 < x
disc 0 {| pos := x; cond_pos := H2 |} =_D f0 x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x0 y : R => Rabs y < x0:R -> R -> Prop
H0:forall x0 : R, (exists y : R, g x0 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H2:0 = x \/ 0 > x
open_set (f0 x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x0 y : R => Rabs y < x0:R -> R -> Prop
H0:forall x0 : R, (exists y : R, g x0 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H2:0 < x

disc 0 {| pos := x; cond_pos := H2 |} =_D f0 x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x0 y : R => Rabs y < x0:R -> R -> Prop
H0:forall x0 : R, (exists y : R, g x0 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H2:0 = x \/ 0 > x
open_set (f0 x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x0 y : R => Rabs y < x0:R -> R -> Prop
H0:forall x0 : R, (exists y : R, g x0 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H2:0 < x

included (fun y : R => Rabs (y - 0) < {| pos := x; cond_pos := H2 |}) (fun y : R => Rabs y < x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x0 y : R => Rabs y < x0:R -> R -> Prop
H0:forall x0 : R, (exists y : R, g x0 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H2:0 < x
included (fun y : R => Rabs y < x) (fun y : R => Rabs (y - 0) < {| pos := x; cond_pos := H2 |})
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x0 y : R => Rabs y < x0:R -> R -> Prop
H0:forall x0 : R, (exists y : R, g x0 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H2:0 = x \/ 0 > x
open_set (f0 x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x0 y : R => Rabs y < x0:R -> R -> Prop
H0:forall x0 : R, (exists y : R, g x0 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H2:0 < x

included (fun y : R => Rabs y < x) (fun y : R => Rabs (y - 0) < {| pos := x; cond_pos := H2 |})
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x0 y : R => Rabs y < x0:R -> R -> Prop
H0:forall x0 : R, (exists y : R, g x0 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H2:0 = x \/ 0 > x
open_set (f0 x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x0 y : R => Rabs y < x0:R -> R -> Prop
H0:forall x0 : R, (exists y : R, g x0 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H2:0 = x \/ 0 > x

open_set (f0 x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x0 y : R => Rabs y < x0:R -> R -> Prop
H0:forall x0 : R, (exists y : R, g x0 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H2:0 = x \/ 0 > x

open_set (fun _ : R => False)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x0 y : R => Rabs y < x0:R -> R -> Prop
H0:forall x0 : R, (exists y : R, g x0 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H2:0 = x \/ 0 > x
(fun _ : R => False) =_D f0 x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x0 y : R => Rabs y < x0:R -> R -> Prop
H0:forall x0 : R, (exists y : R, g x0 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H2:0 = x \/ 0 > x

(fun _ : R => False) =_D f0 x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x0 y : R => Rabs y < x0:R -> R -> Prop
H0:forall x0 : R, (exists y : R, g x0 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H2:0 = x \/ 0 > x

included (fun _ : R => False) (f0 x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x0 y : R => Rabs y < x0:R -> R -> Prop
H0:forall x0 : R, (exists y : R, g x0 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H2:0 = x \/ 0 > x
included (f0 x) (fun _ : R => False)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
D:=fun _ : R => True:R -> Prop
g:=fun x0 y : R => Rabs y < x0:R -> R -> Prop
H0:forall x0 : R, (exists y : R, g x0 y) -> True
f0:={| ind := D; f := g; cond_fam := H0 |}:family
H1:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H2:0 = x \/ 0 > x

included (f0 x) (fun _ : R => False)
unfold included, f0; simpl; unfold g; intros; elim H2; intro; [ rewrite <- H4 in H3; assert (H5 := Rabs_pos x0); elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H5 H3)) | assert (H6 := Rabs_pos x0); assert (H7 := Rlt_trans _ _ _ H3 H4); elim (Rlt_irrefl _ (Rle_lt_trans _ _ _ H6 H7)) ]. Qed. (**********)

forall X : R -> Prop, compact X -> closed_set X

forall X : R -> Prop, compact X -> closed_set X
X:R -> Prop
H:compact X

X =_D adherence X
X:R -> Prop
H:compact X

included X (adherence X)
X:R -> Prop
H:compact X
included (adherence X) X
X:R -> Prop
H:compact X

included (adherence X) X
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:X x

X x
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
X x
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x

X x
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x

(forall y : R, X y -> 0 < Rabs (y - x) / 2) -> X x
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop

(forall x0 : R, (exists y : R, g x0 y) -> D x0) -> D x
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)

covering_open_set X f0 -> D x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering_finite X (subfamily f0 D')

D x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R

0 < alp -> D x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
0 < alp
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y

neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> D x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y
neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
0 < alp
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1
H1:~ X x
H2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2
D:=X:R -> Prop
g:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> Prop
H3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1
H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x
y:R
H12:Rabs (y - x) < alp
H13:X y
y0:R
H15:D' y0
H14:Rabs (y0 - y) < Rabs (y0 - x) / 2
H16:D y0

alp <= Rabs (y0 - x) / 2 -> D x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1
H1:~ X x
H2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2
D:=X:R -> Prop
g:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> Prop
H3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1
H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x
y:R
H12:Rabs (y - x) < alp
H13:X y
y0:R
H15:D' y0
H14:Rabs (y0 - y) < Rabs (y0 - x) / 2
H16:D y0
alp <= Rabs (y0 - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y
neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
0 < alp
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1
H1:~ X x
H2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2
D:=X:R -> Prop
g:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> Prop
H3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1
H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x
y:R
H12:Rabs (y - x) < alp
H13:X y
y0:R
H15:D' y0
H14:Rabs (y0 - y) < Rabs (y0 - x) / 2
H16:D y0
H17:alp <= Rabs (y0 - x) / 2
H18:Rabs (y - x) < Rabs (y0 - x) / 2

Rabs (y0 - x) < Rabs (y0 - x) -> D x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1
H1:~ X x
H2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2
D:=X:R -> Prop
g:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> Prop
H3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1
H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x
y:R
H12:Rabs (y - x) < alp
H13:X y
y0:R
H15:D' y0
H14:Rabs (y0 - y) < Rabs (y0 - x) / 2
H16:D y0
H17:alp <= Rabs (y0 - x) / 2
H18:Rabs (y - x) < Rabs (y0 - x) / 2
Rabs (y0 - x) < Rabs (y0 - x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1
H1:~ X x
H2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2
D:=X:R -> Prop
g:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> Prop
H3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1
H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x
y:R
H12:Rabs (y - x) < alp
H13:X y
y0:R
H15:D' y0
H14:Rabs (y0 - y) < Rabs (y0 - x) / 2
H16:D y0
alp <= Rabs (y0 - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y
neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
0 < alp
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1
H1:~ X x
H2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2
D:=X:R -> Prop
g:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> Prop
H3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1
H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x
y:R
H12:Rabs (y - x) < alp
H13:X y
y0:R
H15:D' y0
H14:Rabs (y0 - y) < Rabs (y0 - x) / 2
H16:D y0
H17:alp <= Rabs (y0 - x) / 2
H18:Rabs (y - x) < Rabs (y0 - x) / 2

Rabs (y0 - x) < Rabs (y0 - x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1
H1:~ X x
H2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2
D:=X:R -> Prop
g:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> Prop
H3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1
H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x
y:R
H12:Rabs (y - x) < alp
H13:X y
y0:R
H15:D' y0
H14:Rabs (y0 - y) < Rabs (y0 - x) / 2
H16:D y0
alp <= Rabs (y0 - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y
neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
0 < alp
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1
H1:~ X x
H2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2
D:=X:R -> Prop
g:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> Prop
H3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1
H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x
y:R
H12:Rabs (y - x) < alp
H13:X y
y0:R
H15:D' y0
H14:Rabs (y0 - y) < Rabs (y0 - x) / 2
H16:D y0
H17:alp <= Rabs (y0 - x) / 2
H18:Rabs (y - x) < Rabs (y0 - x) / 2

Rabs (y0 - x) <= Rabs (y0 - y) + Rabs (y - x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1
H1:~ X x
H2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2
D:=X:R -> Prop
g:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> Prop
H3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1
H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x
y:R
H12:Rabs (y - x) < alp
H13:X y
y0:R
H15:D' y0
H14:Rabs (y0 - y) < Rabs (y0 - x) / 2
H16:D y0
H17:alp <= Rabs (y0 - x) / 2
H18:Rabs (y - x) < Rabs (y0 - x) / 2
Rabs (y0 - y) + Rabs (y - x) < Rabs (y0 - x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1
H1:~ X x
H2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2
D:=X:R -> Prop
g:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> Prop
H3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1
H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x
y:R
H12:Rabs (y - x) < alp
H13:X y
y0:R
H15:D' y0
H14:Rabs (y0 - y) < Rabs (y0 - x) / 2
H16:D y0
alp <= Rabs (y0 - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y
neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
0 < alp
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1
H1:~ X x
H2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2
D:=X:R -> Prop
g:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> Prop
H3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1
H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x
y:R
H12:Rabs (y - x) < alp
H13:X y
y0:R
H15:D' y0
H14:Rabs (y0 - y) < Rabs (y0 - x) / 2
H16:D y0
H17:alp <= Rabs (y0 - x) / 2
H18:Rabs (y - x) < Rabs (y0 - x) / 2

Rabs (y0 - y) + Rabs (y - x) < Rabs (y0 - x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1
H1:~ X x
H2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2
D:=X:R -> Prop
g:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> Prop
H3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1
H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x
y:R
H12:Rabs (y - x) < alp
H13:X y
y0:R
H15:D' y0
H14:Rabs (y0 - y) < Rabs (y0 - x) / 2
H16:D y0
alp <= Rabs (y0 - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y
neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
0 < alp
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y1 : R, intersection_domain V X y1
H1:~ X x
H2:forall y1 : R, X y1 -> 0 < Rabs (y1 - x) / 2
D:=X:R -> Prop
g:=fun y1 z : R => Rabs (y1 - z) < Rabs (y1 - x) / 2 /\ D y1:R -> R -> Prop
H3:forall x0 : R, (exists y1 : R, g x0 y1) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y1 : R, g y1 x0 /\ D' y1
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y1 : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y1
H11:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x
y:R
H12:Rabs (y - x) < alp
H13:X y
y0:R
H15:D' y0
H14:Rabs (y0 - y) < Rabs (y0 - x) / 2
H16:D y0

alp <= Rabs (y0 - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y
neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
0 < alp
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y

neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
0 < alp
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
H9:0 < alp
H10:neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x -> exists y : R, intersection_domain (disc x {| pos := alp; cond_pos := H9 |}) X y
H11:forall x0 : R, disc x {| pos := alp; cond_pos := H9 |} x0 -> neighbourhood (disc x {| pos := alp; cond_pos := H9 |}) x0

disc x {| pos := alp; cond_pos := H9 |} x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R
0 < alp
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
H5:covering_open_set X f0
D':R -> Prop
H6:covering X (subfamily f0 D') /\ family_finite (subfamily f0 D')
H7:forall x0 : R, X x0 -> exists y : R, g y x0 /\ D' y
l:Rlist
H8:forall x0 : R, intersection_domain D D' x0 <-> In x0 l
alp:=MinRlist (AbsList l x):R

0 < alp
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
covering_open_set X f0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)

covering_open_set X f0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)

covering X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
family_open_set f0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:X x0

Rabs (x0 - x0) < Rabs (x0 - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:X x0
D x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
family_open_set f0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:X x0

D x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
family_open_set f0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)

family_open_set f0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:D x0

open_set (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:~ D x0
open_set (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:D x0

open_set (disc x0 {| pos := Rabs (x0 - x) / 2; cond_pos := H2 x0 H5 |})
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:D x0
disc x0 {| pos := Rabs (x0 - x) / 2; cond_pos := H2 x0 H5 |} =_D (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:~ D x0
open_set (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:D x0

disc x0 {| pos := Rabs (x0 - x) / 2; cond_pos := H2 x0 H5 |} =_D (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:~ D x0
open_set (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:D x0

included (disc x0 {| pos := Rabs (x0 - x) / 2; cond_pos := H2 x0 H5 |}) (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:D x0
included (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0) (disc x0 {| pos := Rabs (x0 - x) / 2; cond_pos := H2 x0 H5 |})
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:~ D x0
open_set (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x2 : R, (exists y : R, g x2 y) -> D x2
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:D x0
x1:R
H6:Rabs (x1 - x0) < Rabs (x0 - x) / 2

Rabs (x0 - x1) < Rabs (x0 - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x2 : R, (exists y : R, g x2 y) -> D x2
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:D x0
x1:R
H6:Rabs (x1 - x0) < Rabs (x0 - x) / 2
D x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:D x0
included (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0) (disc x0 {| pos := Rabs (x0 - x) / 2; cond_pos := H2 x0 H5 |})
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:~ D x0
open_set (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x2 : R, (exists y : R, g x2 y) -> D x2
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:D x0
x1:R
H6:Rabs (x1 - x0) < Rabs (x0 - x) / 2

D x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:D x0
included (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0) (disc x0 {| pos := Rabs (x0 - x) / 2; cond_pos := H2 x0 H5 |})
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:~ D x0
open_set (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:D x0

included (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0) (disc x0 {| pos := Rabs (x0 - x) / 2; cond_pos := H2 x0 H5 |})
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:~ D x0
open_set (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:~ D x0

open_set (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:~ D x0

open_set (fun _ : R => False)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:~ D x0
(fun _ : R => False) =_D (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:~ D x0

(fun _ : R => False) =_D (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:~ D x0

included (fun _ : R => False) (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:~ D x0
included (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0) (fun _ : R => False)
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
H3:forall x1 : R, (exists y : R, g x1 y) -> D x1
f0:={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x0:R
H5:~ D x0

included (fun z : R => Rabs (x0 - z) < Rabs (x0 - x) / 2 /\ D x0) (fun _ : R => False)
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop
forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D0 : R -> Prop, covering_finite X (subfamily f0 D0)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
H2:forall y : R, X y -> 0 < Rabs (y - x) / 2
D:=X:R -> Prop
g:=fun y z : R => Rabs (y - z) < Rabs (y - x) / 2 /\ D y:R -> R -> Prop

forall x0 : R, (exists y : R, g x0 y) -> D x0
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x
forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y : R, intersection_domain V X y
H1:~ X x

forall y : R, X y -> 0 < Rabs (y - x) / 2
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y0 : R, intersection_domain V X y0
H1:~ X x
y:R
H2:X y

0 < Rabs (y - x)
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y0 : R, intersection_domain V X y0
H1:~ X x
y:R
H2:X y
0 < / 2
X:R -> Prop
H:forall f0 : family, covering_open_set X f0 -> exists D : R -> Prop, covering_finite X (subfamily f0 D)
x:R
H0:forall V : R -> Prop, neighbourhood V x -> exists y0 : R, intersection_domain V X y0
H1:~ X x
y:R
H2:X y

0 < / 2
apply Rinv_0_lt_compat; prove_sup0. Qed. (**********)

compact (fun _ : R => False)

compact (fun _ : R => False)
f0:family
H:covering_open_set (fun _ : R => False) f0

covering (fun _ : R => False) (subfamily f0 (fun _ : R => False))
f0:family
H:covering_open_set (fun _ : R => False) f0
family_finite (subfamily f0 (fun _ : R => False))
f0:family
H:covering_open_set (fun _ : R => False) f0

family_finite (subfamily f0 (fun _ : R => False))
f0:family
H:covering_open_set (fun _ : R => False) f0
x:R

ind (subfamily f0 (fun _ : R => False)) x <-> In x nil
f0:family
H:covering_open_set (fun _ : R => False) f0
x:R

ind (subfamily f0 (fun _ : R => False)) x -> In x nil
f0:family
H:covering_open_set (fun _ : R => False) f0
x:R
In x nil -> ind (subfamily f0 (fun _ : R => False)) x
f0:family
H:covering_open_set (fun _ : R => False) f0
x:R
H0:ind f0 x /\ False

ind f0 x -> False -> False
f0:family
H:covering_open_set (fun _ : R => False) f0
x:R
In x nil -> ind (subfamily f0 (fun _ : R => False)) x
f0:family
H:covering_open_set (fun _ : R => False) f0
x:R

In x nil -> ind (subfamily f0 (fun _ : R => False)) x
simpl; intro; elim H0. Qed.

forall X1 X2 : R -> Prop, compact X1 -> X1 =_D X2 -> compact X2

forall X1 X2 : R -> Prop, compact X1 -> X1 =_D X2 -> compact X2
X1, X2:R -> Prop
H:forall f1 : family, covering_open_set X1 f1 -> exists D : R -> Prop, covering_finite X1 (subfamily f1 D)
f0:family
H1:covering_open_set X2 f0
H0:forall x : R, X1 x -> X2 x
H2:forall x : R, X2 x -> X1 x

covering_open_set X1 f0
X1, X2:R -> Prop
H:forall f1 : family, covering_open_set X1 f1 -> exists D : R -> Prop, covering_finite X1 (subfamily f1 D)
f0:family
H1:covering_open_set X2 f0
H0:forall x : R, X1 x -> X2 x
H2:forall x : R, X2 x -> X1 x
H3:covering_open_set X1 f0
exists D : R -> Prop, covering_finite X2 (subfamily f0 D)
X1, X2:R -> Prop
H:forall f1 : family, covering_open_set X1 f1 -> exists D : R -> Prop, covering_finite X1 (subfamily f1 D)
f0:family
H0:forall x : R, X1 x -> X2 x
H2:forall x : R, X2 x -> X1 x
H1:covering X2 f0
H3:family_open_set f0

covering X1 f0
X1, X2:R -> Prop
H:forall f1 : family, covering_open_set X1 f1 -> exists D : R -> Prop, covering_finite X1 (subfamily f1 D)
f0:family
H0:forall x : R, X1 x -> X2 x
H2:forall x : R, X2 x -> X1 x
H1:covering X2 f0
H3:family_open_set f0
family_open_set f0
X1, X2:R -> Prop
H:forall f1 : family, covering_open_set X1 f1 -> exists D : R -> Prop, covering_finite X1 (subfamily f1 D)
f0:family
H1:covering_open_set X2 f0
H0:forall x : R, X1 x -> X2 x
H2:forall x : R, X2 x -> X1 x
H3:covering_open_set X1 f0
exists D : R -> Prop, covering_finite X2 (subfamily f0 D)
X1, X2:R -> Prop
H:forall f1 : family, covering_open_set X1 f1 -> exists D : R -> Prop, covering_finite X1 (subfamily f1 D)
f0:family
H0:forall x : R, X1 x -> X2 x
H2:forall x : R, X2 x -> X1 x
H1:covering X2 f0
H3:family_open_set f0

family_open_set f0
X1, X2:R -> Prop
H:forall f1 : family, covering_open_set X1 f1 -> exists D : R -> Prop, covering_finite X1 (subfamily f1 D)
f0:family
H1:covering_open_set X2 f0
H0:forall x : R, X1 x -> X2 x
H2:forall x : R, X2 x -> X1 x
H3:covering_open_set X1 f0
exists D : R -> Prop, covering_finite X2 (subfamily f0 D)
X1, X2:R -> Prop
H:forall f1 : family, covering_open_set X1 f1 -> exists D : R -> Prop, covering_finite X1 (subfamily f1 D)
f0:family
H1:covering_open_set X2 f0
H0:forall x : R, X1 x -> X2 x
H2:forall x : R, X2 x -> X1 x
H3:covering_open_set X1 f0

exists D : R -> Prop, covering_finite X2 (subfamily f0 D)
X1, X2:R -> Prop
H:forall f1 : family, covering_open_set X1 f1 -> exists D0 : R -> Prop, covering_finite X1 (subfamily f1 D0)
f0:family
H1:covering_open_set X2 f0
H0:forall x : R, X1 x -> X2 x
H2:forall x : R, X2 x -> X1 x
H3:covering_open_set X1 f0
D:R -> Prop
H4:covering X1 (subfamily f0 D) /\ family_finite (subfamily f0 D)
H5:covering X1 (subfamily f0 D)
H6:family_finite (subfamily f0 D)

covering X2 (subfamily f0 D)
X1, X2:R -> Prop
H:forall f1 : family, covering_open_set X1 f1 -> exists D0 : R -> Prop, covering_finite X1 (subfamily f1 D0)
f0:family
H1:covering_open_set X2 f0
H0:forall x : R, X1 x -> X2 x
H2:forall x : R, X2 x -> X1 x
H3:covering_open_set X1 f0
D:R -> Prop
H4:covering X1 (subfamily f0 D) /\ family_finite (subfamily f0 D)
H5:covering X1 (subfamily f0 D)
H6:family_finite (subfamily f0 D)
family_finite (subfamily f0 D)
X1, X2:R -> Prop
H:forall f1 : family, covering_open_set X1 f1 -> exists D0 : R -> Prop, covering_finite X1 (subfamily f1 D0)
f0:family
H1:covering_open_set X2 f0
H0:forall x : R, X1 x -> X2 x
H2:forall x : R, X2 x -> X1 x
H3:covering_open_set X1 f0
D:R -> Prop
H4:covering X1 (subfamily f0 D) /\ family_finite (subfamily f0 D)
H5:covering X1 (subfamily f0 D)
H6:family_finite (subfamily f0 D)

family_finite (subfamily f0 D)
apply H6. Qed.
Borel-Lebesgue's lemma

forall a b : R, compact (fun c : R => a <= c <= b)

forall a b : R, compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b

compact (fun c : R => a <= c <= b)
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop

exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a

exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A

exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0

exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)

exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b

exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m

exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)

exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x

exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x

exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop

covering (fun c : R => a <= c <= b) (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hle':x0 <= x

exists y : R, subfamily f0 Db y x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hnle':~ x0 <= x
exists y : R, subfamily f0 Db y x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hle':x0 <= x
H15:a <= x0 <= x

exists y : R, subfamily f0 Db y x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hle':x0 <= x
a <= x0 <= x
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hnle':~ x0 <= x
exists y : R, subfamily f0 Db y x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hle':x0 <= x

a <= x0 <= x
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hnle':~ x0 <= x
exists y : R, subfamily f0 Db y x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hnle':~ x0 <= x

exists y : R, subfamily f0 Db y x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hnle':~ x0 <= x

f0 y0 x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hnle':~ x0 <= x

b - x0 < eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hnle':~ x0 <= x
b - x0 >= 0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hnle':~ x0 <= x

b - x0 < b - x
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hnle':~ x0 <= x
b - x < eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hnle':~ x0 <= x
b - x0 >= 0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hnle':~ x0 <= x

b - x < eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hnle':~ x0 <= x
b - x0 >= 0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hnle':~ x0 <= x

b - x0 >= 0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= b
Hnle':~ x0 <= x

Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop

family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:b - eps < x
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop

exists l : Rlist, forall x0 : R, ind (subfamily f0 Db) x0 <-> In x0 l
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R

ind (subfamily f0 Db) x0 -> In x0 (cons y0 l)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
In x0 (cons y0 l) -> ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
x0:R
H13:ind (subfamily f0 Dx) x0 -> In x0 l
H15:In x0 l -> ind (subfamily f0 Dx) x0
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
H14:ind f0 x0 /\ Db x0
H16:x0 = y0

In x0 (cons y0 l)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
x0:R
H13:ind (subfamily f0 Dx) x0 -> In x0 l
H15:In x0 l -> ind (subfamily f0 Dx) x0
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
H14:ind f0 x0 /\ Db x0
H16:x0 <> y0
In x0 (cons y0 l)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
In x0 (cons y0 l) -> ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
x0:R
H13:ind (subfamily f0 Dx) x0 -> In x0 l
H15:In x0 l -> ind (subfamily f0 Dx) x0
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
H14:ind f0 x0 /\ Db x0
H16:x0 <> y0

In x0 (cons y0 l)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
In x0 (cons y0 l) -> ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
x0:R
H13:ind (subfamily f0 Dx) x0 -> In x0 l
H15:In x0 l -> ind (subfamily f0 Dx) x0
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
H14:ind f0 x0 /\ Db x0
H16:x0 <> y0

ind (subfamily f0 Dx) x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
In x0 (cons y0 l) -> ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
x0:R
H13:ind (subfamily f0 Dx) x0 -> In x0 l
H15:In x0 l -> ind (subfamily f0 Dx) x0
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
H14:ind f0 x0 /\ (Dx x0 \/ x0 = y0)
H16:x0 <> y0
H7:ind f0 x0
H11:Dx x0

ind f0 x0 /\ Dx x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
x0:R
H13:ind (subfamily f0 Dx) x0 -> In x0 l
H15:In x0 l -> ind (subfamily f0 Dx) x0
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
H14:ind f0 x0 /\ (Dx x0 \/ x0 = y0)
H16:x0 <> y0
H7:ind f0 x0
H11:x0 = y0
ind f0 x0 /\ Dx x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
In x0 (cons y0 l) -> ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
x0:R
H13:ind (subfamily f0 Dx) x0 -> In x0 l
H15:In x0 l -> ind (subfamily f0 Dx) x0
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
H14:ind f0 x0 /\ (Dx x0 \/ x0 = y0)
H16:x0 <> y0
H7:ind f0 x0
H11:x0 = y0

ind f0 x0 /\ Dx x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
In x0 (cons y0 l) -> ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R

In x0 (cons y0 l) -> ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:x0 = y0

ind f0 x0 /\ Db x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:In x0 l
ind f0 x0 /\ Db x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:x0 = y0

ind f0 x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:x0 = y0
Db x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:In x0 l
ind f0 x0 /\ Db x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:x0 = y0

Db x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:In x0 l
ind f0 x0 /\ Db x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:In x0 l

ind f0 x0 /\ Db x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:In x0 l

(ind (subfamily f0 Dx) x0 -> In x0 l) -> (In x0 l -> ind (subfamily f0 Dx) x0) -> ind f0 x0 /\ Db x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:In x0 l
H16:In x0 l -> ind (subfamily f0 Dx) x0
H17:ind f0 x0 /\ Dx x0

ind f0 x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:In x0 l
H16:In x0 l -> ind (subfamily f0 Dx) x0
H17:ind f0 x0 /\ Dx x0
Db x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
H4:a <= b <= b
H3:is_upper_bound A b /\ (forall b0 : R, is_upper_bound A b0 -> b <= b0)
y0:R
H6:f0 y0 b
eps:posreal
H8:included (disc b eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:b - eps < x
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:In x0 l
H16:In x0 l -> ind (subfamily f0 Dx) x0
H17:ind f0 x0 /\ Dx x0

Db x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b

exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R

exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
H7:A m'

exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
A m'
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H14:forall x0 : R, A x0 -> x0 <= m
H15:forall b0 : R, is_upper_bound A b0 -> m <= b0
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
H7:A m'

exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
A m'
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H14:forall x0 : R, A x0 -> x0 <= m
H15:forall b0 : R, is_upper_bound A b0 -> m <= b0
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
H7:A m'
H16:m' <= m

exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
A m'
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H14:forall x0 : R, A x0 -> x0 <= m
H15:forall b0 : R, is_upper_bound A b0 -> m <= b0
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
H7:A m'
H16:m' <= m
H17:m < m'

exists D : R -> Prop, covering_finite (fun c : R => a <= c <= b) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H14:forall x0 : R, A x0 -> x0 <= m
H15:forall b0 : R, is_upper_bound A b0 -> m <= b0
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
H7:A m'
H16:m' <= m
m < m'
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
A m'
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H14:forall x0 : R, A x0 -> x0 <= m
H15:forall b0 : R, is_upper_bound A b0 -> m <= b0
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
H7:A m'
H16:m' <= m

m < m'
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
A m'
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H14:forall x0 : R, A x0 -> x0 <= m
H15:forall b0 : R, is_upper_bound A b0 -> m <= b0
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
H7:A m'
H16:m' <= m
Hle':m + eps / 2 <= b

m < m + eps / 2
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H14:forall x0 : R, A x0 -> x0 <= m
H15:forall b0 : R, is_upper_bound A b0 -> m <= b0
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
H7:A m'
H16:m' <= m
Hnle':~ m + eps / 2 <= b
m < b
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
A m'
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H14:forall x0 : R, A x0 -> x0 <= m
H15:forall b0 : R, is_upper_bound A b0 -> m <= b0
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
H7:A m'
H16:m' <= m
Hnle':~ m + eps / 2 <= b

m < b
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
A m'
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H14:forall x0 : R, A x0 -> x0 <= m
H15:forall b0 : R, is_upper_bound A b0 -> m <= b0
H3:m < b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
H7:A m'
H16:m' <= m
Hnle':~ m + eps / 2 <= b

m < b
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H14:forall x0 : R, A x0 -> x0 <= m
H15:forall b0 : R, is_upper_bound A b0 -> m <= b0
H3:m = b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
H7:A m'
H16:m' <= m
Hnle':~ m + eps / 2 <= b
m < b
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
A m'
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H14:forall x0 : R, A x0 -> x0 <= m
H15:forall b0 : R, is_upper_bound A b0 -> m <= b0
H3:m = b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
H7:A m'
H16:m' <= m
Hnle':~ m + eps / 2 <= b

m < b
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
A m'
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R

A m'
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R

a <= m' <= b
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= m') (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R

a <= m'
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
m' <= b
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= m') (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R

a <= m
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
m <= m'
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
m' <= b
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= m') (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R

m <= m'
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
m' <= b
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= m') (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
r:m + eps / 2 <= b

m <= m + eps / 2
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
n:~ m + eps / 2 <= b
m <= b
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
m' <= b
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= m') (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
n:~ m + eps / 2 <= b

m <= b
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
m' <= b
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= m') (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m
H7:m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
n:~ m + eps / 2 <= b

m <= b
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
m' <= b
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= m') (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R

m' <= b
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= m') (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R

exists D : R -> Prop, covering_finite (fun c : R => a <= c <= m') (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop

covering (fun c : R => a <= c <= m') (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hle':x0 <= x

exists y : R, subfamily f0 Db y x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
exists y : R, subfamily f0 Db y x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hle':x0 <= x
H15:a <= x0 <= x

exists y : R, subfamily f0 Db y x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hle':x0 <= x
a <= x0 <= x
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
exists y : R, subfamily f0 Db y x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hle':x0 <= x

a <= x0 <= x
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
exists y : R, subfamily f0 Db y x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x

exists y : R, subfamily f0 Db y x0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x

f0 y0 x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hlt:x0 - m < 0

- (x0 - m) < eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0
x0 - m < eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hlt:x0 - m < 0

m - x0 < m - x
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hlt:x0 - m < 0
m - x < eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0
x0 - m < eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hlt:x0 - m < 0

m - x < eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0
x0 - m < eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hlt:x0 - m < 0

m - eps < x - eps + eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hlt:x0 - m < 0
m - eps = x - eps + (m - x)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0
x0 - m < eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hlt:x0 - m < 0

m - eps < x
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hlt:x0 - m < 0
x = x - eps + eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hlt:x0 - m < 0
m - eps = x - eps + (m - x)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0
x0 - m < eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hlt:x0 - m < 0

x = x - eps + eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hlt:x0 - m < 0
m - eps = x - eps + (m - x)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0
x0 - m < eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hlt:x0 - m < 0

m - eps = x - eps + (m - x)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0
x0 - m < eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0

x0 - m < eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0

x0 - m <= m' - m
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0
m' - m < eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0

m' - m < eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0

m' < m + eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0
m' = m + (m' - m)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0

m' <= m + eps / 2
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0
m + eps / 2 < m + eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0
m' = m + (m' - m)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0

m + eps / 2 < m + eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0
m' = m + (m' - m)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0

0 < 2
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0
2 * (eps / 2) < 2 * eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0
m' = m + (m' - m)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0

2 * (eps / 2) < 2 * eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0
m' = m + (m' - m)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0

1 * eps < 2 * eps
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0
2 <> 0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0
m' = m + (m' - m)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0

2 <> 0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0
m' = m + (m' - m)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Hge:x0 - m >= 0

m' = m + (m' - m)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x
Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:forall x1 : R, a <= x1 <= x -> exists y : R, subfamily f0 Dx y x1
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H14:a <= x0
H18:x0 <= m'
Hnle':~ x0 <= x

Db y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop
family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
H13:family_finite (subfamily f0 Dx)
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x0 : R => Dx x0 \/ x0 = y0:R -> Prop

family_finite (subfamily f0 Db)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R

ind (subfamily f0 Db) x0 -> In x0 (cons y0 l)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
In x0 (cons y0 l) -> ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
x0:R
H13:ind (subfamily f0 Dx) x0 -> In x0 l
H15:In x0 l -> ind (subfamily f0 Dx) x0
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
H14:ind f0 x0 /\ Db x0
Heq:x0 = y0

In x0 (cons y0 l)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
x0:R
H13:ind (subfamily f0 Dx) x0 -> In x0 l
H15:In x0 l -> ind (subfamily f0 Dx) x0
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
H14:ind f0 x0 /\ Db x0
Hneq:x0 <> y0
In x0 (cons y0 l)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
In x0 (cons y0 l) -> ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
x0:R
H13:ind (subfamily f0 Dx) x0 -> In x0 l
H15:In x0 l -> ind (subfamily f0 Dx) x0
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
H14:ind f0 x0 /\ Db x0
Hneq:x0 <> y0

In x0 (cons y0 l)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
In x0 (cons y0 l) -> ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
x0:R
H13:ind (subfamily f0 Dx) x0 -> In x0 l
H15:In x0 l -> ind (subfamily f0 Dx) x0
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
H14:ind f0 x0 /\ (Dx x0 \/ x0 = y0)
Hneq:x0 <> y0
H7:ind f0 x0
H16:Dx x0

ind f0 x0 /\ Dx x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
x0:R
H13:ind (subfamily f0 Dx) x0 -> In x0 l
H15:In x0 l -> ind (subfamily f0 Dx) x0
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
H14:ind f0 x0 /\ (Dx x0 \/ x0 = y0)
Hneq:x0 <> y0
H7:ind f0 x0
H16:x0 = y0
ind f0 x0 /\ Dx x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
In x0 (cons y0 l) -> ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
x0:R
H13:ind (subfamily f0 Dx) x0 -> In x0 l
H15:In x0 l -> ind (subfamily f0 Dx) x0
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
H14:ind f0 x0 /\ (Dx x0 \/ x0 = y0)
Hneq:x0 <> y0
H7:ind f0 x0
H16:x0 = y0

ind f0 x0 /\ Dx x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
In x0 (cons y0 l) -> ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R

In x0 (cons y0 l) -> ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:x0 = y0

ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:In x0 l
ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:x0 = y0

ind f0 x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:x0 = y0
Db x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:In x0 l
ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:x0 = y0

Db x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:In x0 l
ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:In x0 l

ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:In x0 l
H16:In x0 l -> ind (subfamily f0 Dx) x0

ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:In x0 l
H16:In x0 l -> ind (subfamily f0 Dx) x0
H17:ind (subfamily f0 Dx) x0

ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:In x0 l
H16:In x0 l -> ind (subfamily f0 Dx) x0
H17:intersection_domain (ind f0) Dx x0

ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:In x0 l
H16:In x0 l -> ind (subfamily f0 Dx) x0
H17:ind f0 x0 /\ Dx x0

ind (subfamily f0 Db) x0
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:In x0 l
H16:In x0 l -> ind (subfamily f0 Dx) x0
H17:ind f0 x0 /\ Dx x0

ind f0 x0
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:In x0 l
H16:In x0 l -> ind (subfamily f0 Dx) x0
H17:ind f0 x0 /\ Dx x0
Db x0
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x1 : R, a <= x1 <= b -> exists y : R, f0 y x1
H5:forall x1 : R, open_set (f0 x1)
A:=fun x1 : R => a <= x1 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x1) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
x:R
H9:a <= x <= b
Dx:R -> Prop
H12:covering (fun c : R => a <= c <= x) (subfamily f0 Dx)
l:Rlist
H13:forall x1 : R, ind (subfamily f0 Dx) x1 <-> In x1 l
Hltx:m - eps < x
H11:m <> b
m':=Rmin (m + eps / 2) b:R
Db:=fun x1 : R => Dx x1 \/ x1 = y0:R -> Prop
x0:R
H15:In x0 l
H16:In x0 l -> ind (subfamily f0 Dx) x0
H17:ind f0 x0 /\ Dx x0

Db x0
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)

exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:exists x : R, A x /\ m - eps < x <= m

exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:~ (exists x : R, A x /\ m - eps < x <= m)
exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:~ (exists x : R, A x /\ m - eps < x <= m)

exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:~ (exists x : R, A x /\ m - eps < x <= m)
H10:is_upper_bound A m
H11:forall b0 : R, is_upper_bound A b0 -> m <= b0

is_upper_bound A (m - eps) -> exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:~ (exists x : R, A x /\ m - eps < x <= m)
H10:is_upper_bound A m
H11:forall b0 : R, is_upper_bound A b0 -> m <= b0
is_upper_bound A (m - eps)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:~ (exists x : R, A x /\ m - eps < x <= m)
H10:is_upper_bound A m
H11:forall b0 : R, is_upper_bound A b0 -> m <= b0
H12:is_upper_bound A (m - eps)
H13:m <= m - eps

m - eps < m -> exists x : R, A x /\ m - eps < x <= m
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:~ (exists x : R, A x /\ m - eps < x <= m)
H10:is_upper_bound A m
H11:forall b0 : R, is_upper_bound A b0 -> m <= b0
H12:is_upper_bound A (m - eps)
H13:m <= m - eps
m - eps < m
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:~ (exists x : R, A x /\ m - eps < x <= m)
H10:is_upper_bound A m
H11:forall b0 : R, is_upper_bound A b0 -> m <= b0
is_upper_bound A (m - eps)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:~ (exists x : R, A x /\ m - eps < x <= m)
H10:is_upper_bound A m
H11:forall b0 : R, is_upper_bound A b0 -> m <= b0
H12:is_upper_bound A (m - eps)
H13:m <= m - eps

m - eps < m
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:~ (exists x : R, A x /\ m - eps < x <= m)
H10:is_upper_bound A m
H11:forall b0 : R, is_upper_bound A b0 -> m <= b0
is_upper_bound A (m - eps)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:forall x : R, open_set (f0 x)
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:~ (exists x : R, A x /\ m - eps < x <= m)
H10:is_upper_bound A m
H11:forall b0 : R, is_upper_bound A b0 -> m <= b0

is_upper_bound A (m - eps)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)
H10:is_upper_bound A m
H11:forall b0 : R, is_upper_bound A b0 -> m <= b0
P:=fun n : R => A n /\ m - eps < n <= m:R -> Prop
H12:forall n : R, ~ (A n /\ m - eps < n <= m)
x:R
H13:A x
H14:~ A x \/ ~ m - eps < x <= m
H15:~ A x

x <= m - eps
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)
H10:is_upper_bound A m
H11:forall b0 : R, is_upper_bound A b0 -> m <= b0
P:=fun n : R => A n /\ m - eps < n <= m:R -> Prop
H12:forall n : R, ~ (A n /\ m - eps < n <= m)
x:R
H13:A x
H14:~ A x \/ ~ m - eps < x <= m
H15:~ m - eps < x <= m
x <= m - eps
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)
H10:is_upper_bound A m
H11:forall b0 : R, is_upper_bound A b0 -> m <= b0
P:=fun n : R => A n /\ m - eps < n <= m:R -> Prop
H12:forall n : R, ~ (A n /\ m - eps < n <= m)
x:R
H13:A x
H14:~ A x \/ ~ m - eps < x <= m
H15:~ m - eps < x <= m

x <= m - eps
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)
H10:is_upper_bound A m
H11:forall b0 : R, is_upper_bound A b0 -> m <= b0
P:=fun n : R => A n /\ m - eps < n <= m:R -> Prop
H12:forall n : R, ~ (A n /\ m - eps < n <= m)
x:R
H13:A x
H14:~ A x \/ ~ m - eps < x <= m
H15:~ m - eps < x <= m
H16:~ m - eps < x

x <= m - eps
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)
H10:is_upper_bound A m
H11:forall b0 : R, is_upper_bound A b0 -> m <= b0
P:=fun n : R => A n /\ m - eps < n <= m:R -> Prop
H12:forall n : R, ~ (A n /\ m - eps < n <= m)
x:R
H13:A x
H14:~ A x \/ ~ m - eps < x <= m
H15:~ m - eps < x <= m
H16:~ x <= m
x <= m - eps
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)
H10:is_upper_bound A m
H11:forall b0 : R, is_upper_bound A b0 -> m <= b0
P:=fun n : R => A n /\ m - eps < n <= m:R -> Prop
H12:forall n : R, ~ (A n /\ m - eps < n <= m)
x:R
H13:A x
H14:~ A x \/ ~ m - eps < x <= m
H15:~ m - eps < x <= m
H16:~ m - eps < x
H17:x <= m - eps

x <= m - eps
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)
H10:is_upper_bound A m
H11:forall b0 : R, is_upper_bound A b0 -> m <= b0
P:=fun n : R => A n /\ m - eps < n <= m:R -> Prop
H12:forall n : R, ~ (A n /\ m - eps < n <= m)
x:R
H13:A x
H14:~ A x \/ ~ m - eps < x <= m
H15:~ m - eps < x <= m
H16:~ m - eps < x
H17:~ x <= m - eps
x <= m - eps
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)
H10:is_upper_bound A m
H11:forall b0 : R, is_upper_bound A b0 -> m <= b0
P:=fun n : R => A n /\ m - eps < n <= m:R -> Prop
H12:forall n : R, ~ (A n /\ m - eps < n <= m)
x:R
H13:A x
H14:~ A x \/ ~ m - eps < x <= m
H15:~ m - eps < x <= m
H16:~ x <= m
x <= m - eps
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)
H10:is_upper_bound A m
H11:forall b0 : R, is_upper_bound A b0 -> m <= b0
P:=fun n : R => A n /\ m - eps < n <= m:R -> Prop
H12:forall n : R, ~ (A n /\ m - eps < n <= m)
x:R
H13:A x
H14:~ A x \/ ~ m - eps < x <= m
H15:~ m - eps < x <= m
H16:~ m - eps < x
H17:~ x <= m - eps

x <= m - eps
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)
H10:is_upper_bound A m
H11:forall b0 : R, is_upper_bound A b0 -> m <= b0
P:=fun n : R => A n /\ m - eps < n <= m:R -> Prop
H12:forall n : R, ~ (A n /\ m - eps < n <= m)
x:R
H13:A x
H14:~ A x \/ ~ m - eps < x <= m
H15:~ m - eps < x <= m
H16:~ x <= m
x <= m - eps
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:forall x0 : R, open_set (f0 x0)
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
H4:a <= m <= b
y0:R
H6:f0 y0 m
eps:posreal
H8:included (disc m eps) (f0 y0)
H9:~ (exists x0 : R, A x0 /\ m - eps < x0 <= m)
H10:is_upper_bound A m
H11:forall b0 : R, is_upper_bound A b0 -> m <= b0
P:=fun n : R => A n /\ m - eps < n <= m:R -> Prop
H12:forall n : R, ~ (A n /\ m - eps < n <= m)
x:R
H13:A x
H14:~ A x \/ ~ m - eps < x <= m
H15:~ m - eps < x <= m
H16:~ x <= m

x <= m - eps
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)
a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m /\ (forall b0 : R, is_upper_bound A b0 -> m <= b0)

a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:is_upper_bound A m
H4:forall b0 : R, is_upper_bound A b0 -> m <= b0

a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:forall x : R, A x -> x <= m
H4:forall b0 : R, is_upper_bound A b0 -> m <= b0

a <= m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:forall x : R, A x -> x <= m
H4:forall b0 : R, is_upper_bound A b0 -> m <= b0

a <= m
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:forall x : R, A x -> x <= m
H4:forall b0 : R, is_upper_bound A b0 -> m <= b0
m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:forall x : R, A x -> x <= m
H4:forall b0 : R, is_upper_bound A b0 -> m <= b0

m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
H2:exists a0 : R, A a0
m:R
H3:forall x : R, A x -> x <= m
H4:forall b0 : R, is_upper_bound A b0 -> m <= b0

m <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A
exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
H1:bound A

exists a0 : R, A a0
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a
bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H0:A a

bound A
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop

A a
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop

a <= a <= b
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
exists D : R -> Prop, covering_finite (fun c : R => a <= c <= a) (subfamily f0 D)
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:covering (fun c : R => a <= c <= b) f0
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop

exists D : R -> Prop, covering_finite (fun c : R => a <= c <= a) (subfamily f0 D)
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop

a <= a <= b -> exists D : R -> Prop, covering_finite (fun c : R => a <= c <= a) (subfamily f0 D)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
a <= a <= b
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H1:a <= a <= b
y0:R
H2:f0 y0 a
D':=fun x : R => x = y0:R -> Prop

covering (fun c : R => a <= c <= a) (subfamily f0 D')
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H1:a <= a <= b
y0:R
H2:f0 y0 a
D':=fun x : R => x = y0:R -> Prop
family_finite (subfamily f0 D')
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
a <= a <= b
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:family_open_set f0
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H1:a <= a <= b
y0:R
H2:f0 y0 a
D':=fun x0 : R => x0 = y0:R -> Prop
x:R
H3:a <= x <= a

x = a -> exists y : R, f0 y x /\ D' y
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:family_open_set f0
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H1:a <= a <= b
y0:R
H2:f0 y0 a
D':=fun x0 : R => x0 = y0:R -> Prop
x:R
H3:a <= x <= a
x = a
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H1:a <= a <= b
y0:R
H2:f0 y0 a
D':=fun x : R => x = y0:R -> Prop
family_finite (subfamily f0 D')
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
a <= a <= b
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:family_open_set f0
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H1:a <= a <= b
y0:R
H2:f0 y0 a
D':=fun x0 : R => x0 = y0:R -> Prop
x:R
H3:a <= x <= a
H4:x = a

f0 y0 x
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:family_open_set f0
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H1:a <= a <= b
y0:R
H2:f0 y0 a
D':=fun x0 : R => x0 = y0:R -> Prop
x:R
H3:a <= x <= a
H4:x = a
D' y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:family_open_set f0
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H1:a <= a <= b
y0:R
H2:f0 y0 a
D':=fun x0 : R => x0 = y0:R -> Prop
x:R
H3:a <= x <= a
x = a
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H1:a <= a <= b
y0:R
H2:f0 y0 a
D':=fun x : R => x = y0:R -> Prop
family_finite (subfamily f0 D')
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
a <= a <= b
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:family_open_set f0
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H1:a <= a <= b
y0:R
H2:f0 y0 a
D':=fun x0 : R => x0 = y0:R -> Prop
x:R
H3:a <= x <= a
H4:x = a

D' y0
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:family_open_set f0
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H1:a <= a <= b
y0:R
H2:f0 y0 a
D':=fun x0 : R => x0 = y0:R -> Prop
x:R
H3:a <= x <= a
x = a
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H1:a <= a <= b
y0:R
H2:f0 y0 a
D':=fun x : R => x = y0:R -> Prop
family_finite (subfamily f0 D')
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
a <= a <= b
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:family_open_set f0
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H1:a <= a <= b
y0:R
H2:f0 y0 a
D':=fun x0 : R => x0 = y0:R -> Prop
x:R
H3:a <= x <= a

x = a
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H1:a <= a <= b
y0:R
H2:f0 y0 a
D':=fun x : R => x = y0:R -> Prop
family_finite (subfamily f0 D')
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
a <= a <= b
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
H1:a <= a <= b
y0:R
H2:f0 y0 a
D':=fun x : R => x = y0:R -> Prop

family_finite (subfamily f0 D')
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
a <= a <= b
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:family_open_set f0
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H1:a <= a <= b
y0:R
H2:f0 y0 a
D':=fun x0 : R => x0 = y0:R -> Prop
x:R

ind (subfamily f0 D') x -> In x (cons y0 nil)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:family_open_set f0
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H1:a <= a <= b
y0:R
H2:f0 y0 a
D':=fun x0 : R => x0 = y0:R -> Prop
x:R
In x (cons y0 nil) -> ind (subfamily f0 D') x
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
a <= a <= b
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:family_open_set f0
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H1:a <= a <= b
y0:R
H2:f0 y0 a
D':=fun x0 : R => x0 = y0:R -> Prop
x:R
H3:ind f0 x
H4:D' x

x = y0 \/ False
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:family_open_set f0
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H1:a <= a <= b
y0:R
H2:f0 y0 a
D':=fun x0 : R => x0 = y0:R -> Prop
x:R
In x (cons y0 nil) -> ind (subfamily f0 D') x
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
a <= a <= b
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:family_open_set f0
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H1:a <= a <= b
y0:R
H2:f0 y0 a
D':=fun x0 : R => x0 = y0:R -> Prop
x:R

In x (cons y0 nil) -> ind (subfamily f0 D') x
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
a <= a <= b
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x0 : R, a <= x0 <= b -> exists y : R, f0 y x0
H5:family_open_set f0
A:=fun x0 : R => a <= x0 <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x0) (subfamily f0 D)):R -> Prop
H1:a <= a <= b
y0:R
H2:f0 y0 a
D':=fun x0 : R => x0 = y0:R -> Prop
x:R
H4:x = y0

ind f0 x /\ D' x
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop
a <= a <= b
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hle:a <= b
f0:family
H:forall x : R, a <= x <= b -> exists y : R, f0 y x
H5:family_open_set f0
A:=fun x : R => a <= x <= b /\ (exists D : R -> Prop, covering_finite (fun c : R => a <= c <= x) (subfamily f0 D)):R -> Prop

a <= a <= b
a, b:R
Hnle:~ a <= b
compact (fun c : R => a <= c <= b)
a, b:R
Hnle:~ a <= b

compact (fun c : R => a <= c <= b)
a, b:R
Hnle:~ a <= b

compact (fun _ : R => False)
a, b:R
Hnle:~ a <= b
(fun _ : R => False) =_D (fun c : R => a <= c <= b)
a, b:R
Hnle:~ a <= b

(fun _ : R => False) =_D (fun c : R => a <= c <= b)
a, b:R
Hnle:~ a <= b

included (fun _ : R => False) (fun c : R => a <= c <= b)
a, b:R
Hnle:~ a <= b
included (fun c : R => a <= c <= b) (fun _ : R => False)
a, b:R
Hnle:~ a <= b

included (fun c : R => a <= c <= b) (fun _ : R => False)
unfold included; intros; elim H; clear H; intros; assert (H1 := Rle_trans _ _ _ H H0); elim Hnle; apply H1. Qed.

forall X F : R -> Prop, compact X -> closed_set F -> included F X -> compact F

forall X F : R -> Prop, compact X -> closed_set F -> included F X -> compact F
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z

exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop

exists D0 : R -> Prop, covering_finite F (subfamily f0 D0)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop

exists D0 : R -> Prop, covering_finite F (subfamily f0 D0)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop

exists D0 : R -> Prop, covering_finite F (subfamily f0 D0)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop

(forall x : R, (exists y : R, g' x y) -> D' x) -> exists D0 : R -> Prop, covering_finite F (subfamily f0 D0)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family

covering_open_set X f' -> exists D0 : R -> Prop, covering_finite F (subfamily f0 D0)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
covering_open_set X f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:covering_finite X (subfamily f' DX)

covering_finite F (subfamily f0 DX)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
covering_open_set X f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:covering X (subfamily f' DX)
H6:family_finite (subfamily f' DX)

covering F (subfamily f0 DX) /\ family_finite (subfamily f0 DX)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
covering_open_set X f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:covering X (subfamily f' DX)
H6:family_finite (subfamily f' DX)

covering F (subfamily f0 DX)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:covering X (subfamily f' DX)
H6:family_finite (subfamily f' DX)
family_finite (subfamily f0 DX)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
covering_open_set X f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:forall x0 : R, X x0 -> exists y : R, subfamily f' DX y x0
H6:family_finite (subfamily f' DX)
x:R
H7:F x

exists y : R, subfamily f0 DX y x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:covering X (subfamily f' DX)
H6:family_finite (subfamily f' DX)
family_finite (subfamily f0 DX)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
covering_open_set X f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:forall x0 : R, X x0 -> exists y : R, subfamily f' DX y x0
H6:family_finite (subfamily f' DX)
x:R
H7:F x
y0:R
H8:g' y0 x
H9:DX y0

f0 y0 x /\ DX y0
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:covering X (subfamily f' DX)
H6:family_finite (subfamily f' DX)
family_finite (subfamily f0 DX)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
covering_open_set X f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:forall x0 : R, X x0 -> exists y : R, subfamily f' DX y x0
H6:family_finite (subfamily f' DX)
x:R
H7:F x
y0:R
H8:g' y0 x
H9:DX y0

f0 y0 x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:forall x0 : R, X x0 -> exists y : R, subfamily f' DX y x0
H6:family_finite (subfamily f' DX)
x:R
H7:F x
y0:R
H8:g' y0 x
H9:DX y0
DX y0
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:covering X (subfamily f' DX)
H6:family_finite (subfamily f' DX)
family_finite (subfamily f0 DX)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
covering_open_set X f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:forall x0 : R, X x0 -> exists y : R, subfamily f' DX y x0
H6:family_finite (subfamily f' DX)
x:R
H7:F x
y0:R
H8:f0 y0 x \/ complementary F x /\ D y0
H9:DX y0
H10:f0 y0 x

f0 y0 x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:forall x0 : R, X x0 -> exists y : R, subfamily f' DX y x0
H6:family_finite (subfamily f' DX)
x:R
H7:F x
y0:R
H8:f0 y0 x \/ complementary F x /\ D y0
H9:DX y0
H10:complementary F x /\ D y0
f0 y0 x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:forall x0 : R, X x0 -> exists y : R, subfamily f' DX y x0
H6:family_finite (subfamily f' DX)
x:R
H7:F x
y0:R
H8:g' y0 x
H9:DX y0
DX y0
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:covering X (subfamily f' DX)
H6:family_finite (subfamily f' DX)
family_finite (subfamily f0 DX)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
covering_open_set X f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:forall x0 : R, X x0 -> exists y : R, subfamily f' DX y x0
H6:family_finite (subfamily f' DX)
x:R
H7:F x
y0:R
H8:f0 y0 x \/ complementary F x /\ D y0
H9:DX y0
H10:complementary F x /\ D y0

f0 y0 x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:forall x0 : R, X x0 -> exists y : R, subfamily f' DX y x0
H6:family_finite (subfamily f' DX)
x:R
H7:F x
y0:R
H8:g' y0 x
H9:DX y0
DX y0
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:covering X (subfamily f' DX)
H6:family_finite (subfamily f' DX)
family_finite (subfamily f0 DX)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
covering_open_set X f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:forall x0 : R, X x0 -> exists y : R, subfamily f' DX y x0
H6:family_finite (subfamily f' DX)
x:R
H7:F x
y0:R
H8:g' y0 x
H9:DX y0

DX y0
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:covering X (subfamily f' DX)
H6:family_finite (subfamily f' DX)
family_finite (subfamily f0 DX)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
covering_open_set X f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:covering X (subfamily f' DX)
H6:family_finite (subfamily f' DX)

family_finite (subfamily f0 DX)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
covering_open_set X f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:covering X (subfamily f' DX)
l:Rlist
H6:forall x0 : R, ind (subfamily f' DX) x0 <-> In x0 l
x:R
H7:ind (subfamily f' DX) x -> In x l
H8:In x l -> ind (subfamily f' DX) x

ind (subfamily f0 DX) x <-> In x l
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
covering_open_set X f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:covering X (subfamily f' DX)
l:Rlist
H6:forall x0 : R, ind (subfamily f' DX) x0 <-> In x0 l
x:R
H7:ind (subfamily f' DX) x -> In x l
H8:In x l -> ind (subfamily f' DX) x

ind (subfamily f0 DX) x -> In x l
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:covering X (subfamily f' DX)
l:Rlist
H6:forall x0 : R, ind (subfamily f' DX) x0 <-> In x0 l
x:R
H7:ind (subfamily f' DX) x -> In x l
H8:In x l -> ind (subfamily f' DX) x
In x l -> ind (subfamily f0 DX) x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
covering_open_set X f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H4:covering_open_set X f'
DX:R -> Prop
H5:covering X (subfamily f' DX)
l:Rlist
H6:forall x0 : R, ind (subfamily f' DX) x0 <-> In x0 l
x:R
H7:ind (subfamily f' DX) x -> In x l
H8:In x l -> ind (subfamily f' DX) x

In x l -> ind (subfamily f0 DX) x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
covering_open_set X f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family

covering_open_set X f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0

covering X f' /\ family_open_set f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0

covering X f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
family_open_set f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:forall x0 : R, F x0 -> exists y : R, f0 y x0
H4:family_open_set f0
x:R
H5:X x

exists y : R, f' y x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
family_open_set f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:forall x0 : R, F x0 -> exists y : R, f0 y x0
H4:family_open_set f0
x:R
H5:X x
H6:F x

exists y : R, f' y x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:forall x0 : R, F x0 -> exists y : R, f0 y x0
H4:family_open_set f0
x:R
H5:X x
H6:~ F x
exists y : R, f' y x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
family_open_set f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:forall x0 : R, F x0 -> exists y : R, f0 y x0
H4:family_open_set f0
x:R
H5:X x
H6:~ F x

exists y : R, f' y x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
family_open_set f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:forall x0 : R, F x0 -> exists y : R, f0 y x0
H4:family_open_set f0
x:R
H5:X x
H6:~ F x

(exists z : R, D z) -> exists y : R, f' y x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:forall x0 : R, F x0 -> exists y : R, f0 y x0
H4:family_open_set f0
x:R
H5:X x
H6:~ F x
exists z : R, D z
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
family_open_set f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x1 y : R => f0 x1 y \/ complementary F y /\ D x1:R -> R -> Prop
D':=D:R -> Prop
H3:forall x1 : R, (exists y : R, g' x1 y) -> D' x1
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:forall x1 : R, F x1 -> exists y : R, f0 y x1
H4:family_open_set f0
x:R
H5:X x
H6:~ F x
x0:R
H7:D x0

complementary F x /\ D x0
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:forall x0 : R, F x0 -> exists y : R, f0 y x0
H4:family_open_set f0
x:R
H5:X x
H6:~ F x
exists z : R, D z
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
family_open_set f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x1 y : R => f0 x1 y \/ complementary F y /\ D x1:R -> R -> Prop
D':=D:R -> Prop
H3:forall x1 : R, (exists y : R, g' x1 y) -> D' x1
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:forall x1 : R, F x1 -> exists y : R, f0 y x1
H4:family_open_set f0
x:R
H5:X x
H6:~ F x
x0:R
H7:D x0

complementary F x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x1 y : R => f0 x1 y \/ complementary F y /\ D x1:R -> R -> Prop
D':=D:R -> Prop
H3:forall x1 : R, (exists y : R, g' x1 y) -> D' x1
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:forall x1 : R, F x1 -> exists y : R, f0 y x1
H4:family_open_set f0
x:R
H5:X x
H6:~ F x
x0:R
H7:D x0
D x0
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:forall x0 : R, F x0 -> exists y : R, f0 y x0
H4:family_open_set f0
x:R
H5:X x
H6:~ F x
exists z : R, D z
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
family_open_set f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x1 y : R => f0 x1 y \/ complementary F y /\ D x1:R -> R -> Prop
D':=D:R -> Prop
H3:forall x1 : R, (exists y : R, g' x1 y) -> D' x1
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:forall x1 : R, F x1 -> exists y : R, f0 y x1
H4:family_open_set f0
x:R
H5:X x
H6:~ F x
x0:R
H7:D x0

D x0
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:forall x0 : R, F x0 -> exists y : R, f0 y x0
H4:family_open_set f0
x:R
H5:X x
H6:~ F x
exists z : R, D z
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
family_open_set f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:forall x0 : R, F x0 -> exists y : R, f0 y x0
H4:family_open_set f0
x:R
H5:X x
H6:~ F x

exists z : R, D z
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
family_open_set f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:forall x0 : R, F x0 -> exists y : R, f0 y x0
H4:family_open_set f0
x:R
H5:X x
H6:~ F x
z0:R
H7:F z0

exists z : R, D z
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
family_open_set f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:forall x0 : R, F x0 -> exists y : R, f0 y x0
H4:family_open_set f0
x:R
H5:X x
H6:~ F x
z0:R
H7:F z0
H8:exists y : R, f0 y z0

exists z : R, D z
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
family_open_set f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
H3:forall x : R, (exists y : R, g' x y) -> D' x
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0

family_open_set f'
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:D x

open_set (fun y : R => f0 x y \/ complementary F y /\ D x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:~ D x
open_set (fun y : R => f0 x y \/ complementary F y /\ D x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:D x

open_set (union_domain (f0 x) (complementary F))
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:D x
union_domain (f0 x) (complementary F) =_D (fun y : R => f0 x y \/ complementary F y /\ D x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:~ D x
open_set (fun y : R => f0 x y \/ complementary F y /\ D x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:D x

open_set (f0 x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:D x
open_set (complementary F)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:D x
union_domain (f0 x) (complementary F) =_D (fun y : R => f0 x y \/ complementary F y /\ D x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:~ D x
open_set (fun y : R => f0 x y \/ complementary F y /\ D x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:D x

open_set (complementary F)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:D x
union_domain (f0 x) (complementary F) =_D (fun y : R => f0 x y \/ complementary F y /\ D x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:~ D x
open_set (fun y : R => f0 x y \/ complementary F y /\ D x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:D x

union_domain (f0 x) (complementary F) =_D (fun y : R => f0 x y \/ complementary F y /\ D x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:~ D x
open_set (fun y : R => f0 x y \/ complementary F y /\ D x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:D x

included (union_domain (f0 x) (complementary F)) (fun y : R => f0 x y \/ complementary F y /\ D x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:D x
included (fun y : R => f0 x y \/ complementary F y /\ D x) (union_domain (f0 x) (complementary F))
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:~ D x
open_set (fun y : R => f0 x y \/ complementary F y /\ D x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x1 y : R => f0 x1 y \/ complementary F y /\ D x1:R -> R -> Prop
D':=D:R -> Prop
H3:forall x1 : R, (exists y : R, g' x1 y) -> D' x1
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:D x
x0:R
H6:f0 x x0 \/ ~ F x0

f0 x x0 \/ ~ F x0 /\ D x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:D x
included (fun y : R => f0 x y \/ complementary F y /\ D x) (union_domain (f0 x) (complementary F))
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:~ D x
open_set (fun y : R => f0 x y \/ complementary F y /\ D x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:D x

included (fun y : R => f0 x y \/ complementary F y /\ D x) (union_domain (f0 x) (complementary F))
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:~ D x
open_set (fun y : R => f0 x y \/ complementary F y /\ D x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x1 y : R => f0 x1 y \/ complementary F y /\ D x1:R -> R -> Prop
D':=D:R -> Prop
H3:forall x1 : R, (exists y : R, g' x1 y) -> D' x1
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:D x
x0:R
H6:f0 x x0 \/ ~ F x0 /\ D x

f0 x x0 \/ ~ F x0
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:~ D x
open_set (fun y : R => f0 x y \/ complementary F y /\ D x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:~ D x

open_set (fun y : R => f0 x y \/ complementary F y /\ D x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:~ D x

open_set (f0 x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:~ D x
f0 x =_D (fun y : R => f0 x y \/ complementary F y /\ D x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:~ D x

f0 x =_D (fun y : R => f0 x y \/ complementary F y /\ D x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:~ D x

included (f0 x) (fun y : R => f0 x y \/ complementary F y /\ D x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:~ D x
included (fun y : R => f0 x y \/ complementary F y /\ D x) (f0 x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
H3:forall x0 : R, (exists y : R, g' x0 y) -> D' x0
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:~ D x

included (fun y : R => f0 x y \/ complementary F y /\ D x) (f0 x)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x1 y : R => f0 x1 y \/ complementary F y /\ D x1:R -> R -> Prop
D':=D:R -> Prop
H3:forall x1 : R, (exists y : R, g' x1 y) -> D' x1
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:~ D x
x0:R
H6:f0 x x0 \/ ~ F x0 /\ D x

f0 x x0
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x1 y : R => f0 x1 y \/ complementary F y /\ D x1:R -> R -> Prop
D':=D:R -> Prop
H3:forall x1 : R, (exists y : R, g' x1 y) -> D' x1
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:~ D x
x0:R
H6:f0 x x0 \/ ~ F x0 /\ D x
H7:f0 x x0

f0 x x0
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x1 y : R => f0 x1 y \/ complementary F y /\ D x1:R -> R -> Prop
D':=D:R -> Prop
H3:forall x1 : R, (exists y : R, g' x1 y) -> D' x1
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:~ D x
x0:R
H6:f0 x x0 \/ ~ F x0 /\ D x
H7:~ F x0 /\ D x
f0 x x0
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x1 y : R => f0 x1 y \/ complementary F y /\ D x1:R -> R -> Prop
D':=D:R -> Prop
H3:forall x1 : R, (exists y : R, g' x1 y) -> D' x1
f':={| ind := D'; f := g'; cond_fam := H3 |}:family
H2:covering F f0
H4:family_open_set f0
x:R
H5:~ D x
x0:R
H6:f0 x x0 \/ ~ F x0 /\ D x
H7:~ F x0 /\ D x

f0 x x0
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop
forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x y : R => f0 x y \/ complementary F y /\ D x:R -> R -> Prop
D':=D:R -> Prop

forall x : R, (exists y : R, g' x y) -> D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
x:R
H3:exists y : R, g' x y
y0:R
H4:f0 x y0 \/ complementary F y0 /\ D x
H5:f0 x y0

D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
x:R
H3:exists y : R, g' x y
y0:R
H4:f0 x y0 \/ complementary F y0 /\ D x
H5:complementary F y0 /\ D x
D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H0:open_set (complementary F)
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:exists z : R, F z
D:=ind f0:R -> Prop
g:=f f0:R -> R -> Prop
g':=fun x0 y : R => f0 x0 y \/ complementary F y /\ D x0:R -> R -> Prop
D':=D:R -> Prop
x:R
H3:exists y : R, g' x y
y0:R
H4:f0 x y0 \/ complementary F y0 /\ D x
H5:complementary F y0 /\ D x

D' x
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)

exists D : R -> Prop, covering_finite F (subfamily f0 D)
(* Cas ou F est l'ensemble vide *)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)

compact F -> exists D : R -> Prop, covering_finite F (subfamily f0 D)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
compact F
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)

compact F
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)

compact (fun _ : R => False)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
(fun _ : R => False) =_D F
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)

(fun _ : R => False) =_D F
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)

included (fun _ : R => False) F
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)
included F (fun _ : R => False)
X, F:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H0:closed_set F
H1:included F X
f0:family
H2:covering_open_set F f0
Hyp_F_NE:~ (exists z : R, F z)

included F (fun _ : R => False)
assert (H3 := not_ex_all_not _ _ Hyp_F_NE); unfold included; intros; elim (H3 x); apply H4. Qed. (**********)

forall X : R -> Prop, closed_set X -> bounded X -> compact X

forall X : R -> Prop, closed_set X -> bounded X -> compact X
X:R -> Prop
H:closed_set X
H0:exists m M : R, forall x : R, X x -> m <= x <= M

compact X
X:R -> Prop
H:closed_set X
m:R
H0:exists M : R, forall x : R, X x -> m <= x <= M

compact X
X:R -> Prop
H:closed_set X
m, M:R
H0:forall x : R, X x -> m <= x <= M

compact X
X:R -> Prop
H:closed_set X
m, M:R
H0:forall x : R, X x -> m <= x <= M
H1:compact (fun c : R => m <= c <= M)

compact X
apply (compact_P4 (fun c:R => m <= c <= M) X H1 H H0). Qed. (**********)

forall X : R -> Prop, compact X <-> closed_set X /\ bounded X

forall X : R -> Prop, compact X <-> closed_set X /\ bounded X
X:R -> Prop

compact X -> closed_set X /\ bounded X
X:R -> Prop
closed_set X /\ bounded X -> compact X
X:R -> Prop

closed_set X /\ bounded X -> compact X
intro; elim H; clear H; intros; apply (compact_P5 _ H H0). Qed. Definition image_dir (f:R -> R) (D:R -> Prop) (x:R) : Prop := exists y : R, x = f y /\ D y. (**********)

forall (f0 : R -> R) (X : R -> Prop), (forall x : R, continuity_pt f0 x) -> compact X -> compact (image_dir f0 X)

forall (f0 : R -> R) (X : R -> Prop), (forall x : R, continuity_pt f0 x) -> compact X -> compact (image_dir f0 X)
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D : R -> Prop, covering_finite X (subfamily f2 D)
f1:family
H1:covering (image_dir f0 X) f1 /\ family_open_set f1

exists D : R -> Prop, covering_finite (image_dir f0 X) (subfamily f1 D)
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D : R -> Prop, covering_finite X (subfamily f2 D)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1

exists D : R -> Prop, covering_finite (image_dir f0 X) (subfamily f1 D)
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop

exists D0 : R -> Prop, covering_finite (image_dir f0 X) (subfamily f1 D0)
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop

exists D0 : R -> Prop, covering_finite (image_dir f0 X) (subfamily f1 D0)
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop

(forall x : R, (exists y : R, g x y) -> D x) -> exists D0 : R -> Prop, covering_finite (image_dir f0 X) (subfamily f1 D0)
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
forall x : R, (exists y : R, g x y) -> D x
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
H3:forall x : R, (exists y : R, g x y) -> D x
f':={| ind := D; f := g; cond_fam := H3 |}:family

exists D0 : R -> Prop, covering_finite (image_dir f0 X) (subfamily f1 D0)
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
forall x : R, (exists y : R, g x y) -> D x
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
H3:forall x : R, (exists y : R, g x y) -> D x
f':={| ind := D; f := g; cond_fam := H3 |}:family

covering_open_set X f' -> exists D0 : R -> Prop, covering_finite (image_dir f0 X) (subfamily f1 D0)
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
H3:forall x : R, (exists y : R, g x y) -> D x
f':={| ind := D; f := g; cond_fam := H3 |}:family
covering_open_set X f'
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
forall x : R, (exists y : R, g x y) -> D x
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
H3:forall x : R, (exists y : R, g x y) -> D x
f':={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f'
D':R -> Prop
H5:covering_finite X (subfamily f' D')

covering_finite (image_dir f0 X) (subfamily f1 D')
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
H3:forall x : R, (exists y : R, g x y) -> D x
f':={| ind := D; f := g; cond_fam := H3 |}:family
covering_open_set X f'
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
forall x : R, (exists y : R, g x y) -> D x
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
H3:forall x : R, (exists y : R, g x y) -> D x
f':={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f'
D':R -> Prop
H5:covering X (subfamily f' D')
H6:family_finite (subfamily f' D')

covering (image_dir f0 X) (subfamily f1 D')
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
H3:forall x : R, (exists y : R, g x y) -> D x
f':={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f'
D':R -> Prop
H5:covering X (subfamily f' D')
H6:family_finite (subfamily f' D')
family_finite (subfamily f1 D')
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
H3:forall x : R, (exists y : R, g x y) -> D x
f':={| ind := D; f := g; cond_fam := H3 |}:family
covering_open_set X f'
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
forall x : R, (exists y : R, g x y) -> D x
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
H3:forall x : R, (exists y : R, g x y) -> D x
f':={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f'
D':R -> Prop
H5:covering X (subfamily f' D')
H6:family_finite (subfamily f' D')

family_finite (subfamily f1 D')
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
H3:forall x : R, (exists y : R, g x y) -> D x
f':={| ind := D; f := g; cond_fam := H3 |}:family
covering_open_set X f'
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
forall x : R, (exists y : R, g x y) -> D x
f0:R -> R
X:R -> Prop
H:forall x0 : R, continuity_pt f0 x0
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x0 y : R => image_rec f0 (f1 x0) y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f':={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f'
D':R -> Prop
H5:covering X (subfamily f' D')
H6:exists l0 : Rlist, forall x0 : R, ind (subfamily f' D') x0 <-> In x0 l0
l:Rlist
H7:forall x0 : R, ind (subfamily f' D') x0 <-> In x0 l
x:R
H8:ind (subfamily f' D') x -> In x l
H9:In x l -> ind (subfamily f' D') x
H10:ind (subfamily f1 D') x

In x l
f0:R -> R
X:R -> Prop
H:forall x0 : R, continuity_pt f0 x0
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x0 y : R => image_rec f0 (f1 x0) y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f':={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f'
D':R -> Prop
H5:covering X (subfamily f' D')
H6:exists l0 : Rlist, forall x0 : R, ind (subfamily f' D') x0 <-> In x0 l0
l:Rlist
H7:forall x0 : R, ind (subfamily f' D') x0 <-> In x0 l
x:R
H8:ind (subfamily f' D') x -> In x l
H9:In x l -> ind (subfamily f' D') x
H10:In x l
ind (subfamily f1 D') x
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
H3:forall x : R, (exists y : R, g x y) -> D x
f':={| ind := D; f := g; cond_fam := H3 |}:family
covering_open_set X f'
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
forall x : R, (exists y : R, g x y) -> D x
f0:R -> R
X:R -> Prop
H:forall x0 : R, continuity_pt f0 x0
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x0 y : R => image_rec f0 (f1 x0) y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f':={| ind := D; f := g; cond_fam := H3 |}:family
H4:covering_open_set X f'
D':R -> Prop
H5:covering X (subfamily f' D')
H6:exists l0 : Rlist, forall x0 : R, ind (subfamily f' D') x0 <-> In x0 l0
l:Rlist
H7:forall x0 : R, ind (subfamily f' D') x0 <-> In x0 l
x:R
H8:ind (subfamily f' D') x -> In x l
H9:In x l -> ind (subfamily f' D') x
H10:In x l

ind (subfamily f1 D') x
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
H3:forall x : R, (exists y : R, g x y) -> D x
f':={| ind := D; f := g; cond_fam := H3 |}:family
covering_open_set X f'
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
forall x : R, (exists y : R, g x y) -> D x
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
H3:forall x : R, (exists y : R, g x y) -> D x
f':={| ind := D; f := g; cond_fam := H3 |}:family

covering_open_set X f'
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
forall x : R, (exists y : R, g x y) -> D x
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
H3:forall x : R, (exists y : R, g x y) -> D x
f':={| ind := D; f := g; cond_fam := H3 |}:family

covering X f'
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
H3:forall x : R, (exists y : R, g x y) -> D x
f':={| ind := D; f := g; cond_fam := H3 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
forall x : R, (exists y : R, g x y) -> D x
f0:R -> R
X:R -> Prop
H:forall x0 : R, continuity_pt f0 x0
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:forall x0 : R, (exists y : R, x0 = f0 y /\ X y) -> exists y : R, f1 y x0
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x0 y : R => image_rec f0 (f1 x0) y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f':={| ind := D; f := g; cond_fam := H3 |}:family
x:R
H4:X x

exists y : R, f0 x = f0 y /\ X y
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
H3:forall x : R, (exists y : R, g x y) -> D x
f':={| ind := D; f := g; cond_fam := H3 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
forall x : R, (exists y : R, g x y) -> D x
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
H3:forall x : R, (exists y : R, g x y) -> D x
f':={| ind := D; f := g; cond_fam := H3 |}:family

family_open_set f'
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
forall x : R, (exists y : R, g x y) -> D x
f0:R -> R
X:R -> Prop
H:forall x0 : R, continuity_pt f0 x0
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:forall x0 : R, open_set (f1 x0)
D:=ind f1:R -> Prop
g:=fun x0 y : R => image_rec f0 (f1 x0) y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f':={| ind := D; f := g; cond_fam := H3 |}:family
x:R

(fun y : R => image_rec f0 (f1 x) y) = image_rec f0 (f1 x) -> open_set (fun y : R => image_rec f0 (f1 x) y)
f0:R -> R
X:R -> Prop
H:forall x0 : R, continuity_pt f0 x0
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:forall x0 : R, open_set (f1 x0)
D:=ind f1:R -> Prop
g:=fun x0 y : R => image_rec f0 (f1 x0) y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f':={| ind := D; f := g; cond_fam := H3 |}:family
x:R
(fun y : R => image_rec f0 (f1 x) y) = image_rec f0 (f1 x)
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
forall x : R, (exists y : R, g x y) -> D x
f0:R -> R
X:R -> Prop
H:forall x0 : R, continuity_pt f0 x0
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:forall x0 : R, open_set (f1 x0)
D:=ind f1:R -> Prop
g:=fun x0 y : R => image_rec f0 (f1 x0) y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f':={| ind := D; f := g; cond_fam := H3 |}:family
x:R
H4:(fun y : R => image_rec f0 (f1 x) y) = image_rec f0 (f1 x)

open_set (image_rec f0 (f1 x))
f0:R -> R
X:R -> Prop
H:forall x0 : R, continuity_pt f0 x0
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:forall x0 : R, open_set (f1 x0)
D:=ind f1:R -> Prop
g:=fun x0 y : R => image_rec f0 (f1 x0) y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f':={| ind := D; f := g; cond_fam := H3 |}:family
x:R
(fun y : R => image_rec f0 (f1 x) y) = image_rec f0 (f1 x)
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
forall x : R, (exists y : R, g x y) -> D x
f0:R -> R
X:R -> Prop
H:forall x0 : R, continuity_pt f0 x0
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:forall x0 : R, open_set (f1 x0)
D:=ind f1:R -> Prop
g:=fun x0 y : R => image_rec f0 (f1 x0) y:R -> R -> Prop
H3:forall x0 : R, (exists y : R, g x0 y) -> D x0
f':={| ind := D; f := g; cond_fam := H3 |}:family
x:R

(fun y : R => image_rec f0 (f1 x) y) = image_rec f0 (f1 x)
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop
forall x : R, (exists y : R, g x y) -> D x
f0:R -> R
X:R -> Prop
H:forall x : R, continuity_pt f0 x
H0:forall f2 : family, covering_open_set X f2 -> exists D0 : R -> Prop, covering_finite X (subfamily f2 D0)
f1:family
H1:covering (image_dir f0 X) f1
H2:family_open_set f1
D:=ind f1:R -> Prop
g:=fun x y : R => image_rec f0 (f1 x) y:R -> R -> Prop

forall x : R, (exists y : R, g x y) -> D x
intros; apply (cond_fam f1); unfold g in H3; unfold image_rec in H3; elim H3; intros; exists (f0 x0); apply H4. Qed.

forall (f0 : R -> R) (a b : R), a <= b -> (forall c : R, a <= c <= b -> continuity_pt f0 c) -> exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)

forall (f0 : R -> R) (a b : R), a <= b -> (forall c : R, a <= c <= b -> continuity_pt f0 c) -> exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b

exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R

exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R

0 < b - a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a

exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a

continuity h
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x < a

continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x < a
eps:R
H4:eps > 0

a - x > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x < a
eps:R
H4:eps > 0
forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < a - x -> Rabs (h x0 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x < a
eps:R
H4:eps > 0

forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < a - x -> Rabs (h x0 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x < a
eps:R
H4:eps > 0
x0:R
H5:Rabs (x0 - x) < a - x

Rabs ((if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b) - (if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b)) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x < a
eps:R
H4:eps > 0
x0:R
H5:Rabs (x0 - x) < a - x
r:x <= a

Rabs ((if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b) - f0 a) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x < a
eps:R
H4:eps > 0
x0:R
H5:Rabs (x0 - x) < a - x
x <= a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x < a
eps:R
H4:eps > 0
x0:R
H5:Rabs (x0 - x) < a - x
r:x <= a
r0:x0 <= a

Rabs (f0 a - f0 a) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x < a
eps:R
H4:eps > 0
x0:R
H5:Rabs (x0 - x) < a - x
r:x <= a
x0 <= a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x < a
eps:R
H4:eps > 0
x0:R
H5:Rabs (x0 - x) < a - x
x <= a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x < a
eps:R
H4:eps > 0
x0:R
H5:Rabs (x0 - x) < a - x
r:x <= a

x0 <= a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x < a
eps:R
H4:eps > 0
x0:R
H5:Rabs (x0 - x) < a - x
x <= a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x < a
eps:R
H4:eps > 0
x0:R
H5:Rabs (x0 - x) < a - x
r:x <= a

x0 + - x <= Rabs (x0 - x)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x < a
eps:R
H4:eps > 0
x0:R
H5:Rabs (x0 - x) < a - x
r:x <= a
Rabs (x0 - x) < a + - x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x < a
eps:R
H4:eps > 0
x0:R
H5:Rabs (x0 - x) < a - x
x <= a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x < a
eps:R
H4:eps > 0
x0:R
H5:Rabs (x0 - x) < a - x
r:x <= a

Rabs (x0 - x) < a + - x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x < a
eps:R
H4:eps > 0
x0:R
H5:Rabs (x0 - x) < a - x
x <= a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x < a
eps:R
H4:eps > 0
x0:R
H5:Rabs (x0 - x) < a - x

x <= a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a

continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a

continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a

a <= a <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b

continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond a x1 /\ Rabs (x1 - a) < alp -> Rabs (f0 x1 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x1 : R, D_x no_cond a x1 /\ Rabs (x1 - a) < x0 -> Rabs (f0 x1 - f0 a) < eps)

Rmin x0 (b - a) > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond a x1 /\ Rabs (x1 - a) < alp -> Rabs (f0 x1 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x1 : R, D_x no_cond a x1 /\ Rabs (x1 - a) < x0 -> Rabs (f0 x1 - f0 a) < eps)
forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (b - a) -> Rabs (h x1 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond a x1 /\ Rabs (x1 - a) < alp -> Rabs (f0 x1 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x1 : R, D_x no_cond a x1 /\ Rabs (x1 - a) < x0 -> Rabs (f0 x1 - f0 a) < eps)
r:x0 <= b - a

x0 > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond a x1 /\ Rabs (x1 - a) < alp -> Rabs (f0 x1 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x1 : R, D_x no_cond a x1 /\ Rabs (x1 - a) < x0 -> Rabs (f0 x1 - f0 a) < eps)
n:~ x0 <= b - a
b - a > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond a x1 /\ Rabs (x1 - a) < alp -> Rabs (f0 x1 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x1 : R, D_x no_cond a x1 /\ Rabs (x1 - a) < x0 -> Rabs (f0 x1 - f0 a) < eps)
forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (b - a) -> Rabs (h x1 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond a x1 /\ Rabs (x1 - a) < alp -> Rabs (f0 x1 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x1 : R, D_x no_cond a x1 /\ Rabs (x1 - a) < x0 -> Rabs (f0 x1 - f0 a) < eps)
n:~ x0 <= b - a

b - a > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond a x1 /\ Rabs (x1 - a) < alp -> Rabs (f0 x1 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x1 : R, D_x no_cond a x1 /\ Rabs (x1 - a) < x0 -> Rabs (f0 x1 - f0 a) < eps)
forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (b - a) -> Rabs (h x1 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond a x1 /\ Rabs (x1 - a) < alp -> Rabs (f0 x1 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x1 : R, D_x no_cond a x1 /\ Rabs (x1 - a) < x0 -> Rabs (f0 x1 - f0 a) < eps)

forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (b - a) -> Rabs (h x1 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)

x1 < b -> Rabs (h x1 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a

Rabs ((if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b) - f0 a) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
x <= a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hlta:x1 <= a

Rabs (f0 a - f0 a) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a
Rabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 a) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
x <= a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a

Rabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 a) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
x <= a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a
Hleb:x1 <= b

Rabs (f0 x1 - f0 a) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a
x1 <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
x <= a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a
Hleb:x1 <= b
H11:x0 > 0
H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps

D_x no_cond a x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a
Hleb:x1 <= b
H11:x0 > 0
H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps
Rabs (x1 - a) < x0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a
x1 <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
x <= a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a
Hleb:x1 <= b
H11:x0 > 0
H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps

True
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a
Hleb:x1 <= b
H11:x0 > 0
H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps
a <> x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a
Hleb:x1 <= b
H11:x0 > 0
H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps
Rabs (x1 - a) < x0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a
x1 <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
x <= a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a
Hleb:x1 <= b
H11:x0 > 0
H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps

a <> x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a
Hleb:x1 <= b
H11:x0 > 0
H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps
Rabs (x1 - a) < x0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a
x1 <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
x <= a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a
Hleb:x1 <= b
H11:x0 > 0
H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps

Rabs (x1 - a) < x0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a
x1 <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
x <= a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a
Hleb:x1 <= b
H11:x0 > 0
H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps

Rabs (x1 - a) < Rmin x0 (b - a)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a
Hleb:x1 <= b
H11:x0 > 0
H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps
Rmin x0 (b - a) <= x0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a
x1 <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
x <= a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a
Hleb:x1 <= b
H11:x0 > 0
H12:forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps

Rmin x0 (b - a) <= x0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a
x1 <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
x <= a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
r:x <= a
Hnlea:~ x1 <= a

x1 <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b
x <= a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
H10:x1 < b

x <= a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)
x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - x) < Rmin x0 (b - a)

x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - a) < Rmin x0 (b - a)

x1 + - a <= Rabs (x1 - a)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - a) < Rmin x0 (b - a)
Rabs (x1 - a) < b + - a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - a) < Rmin x0 (b - a)

Rabs (x1 - a) < b + - a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - a) < Rmin x0 (b - a)

Rabs (x1 - a) < Rmin x0 (b - a)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - a) < Rmin x0 (b - a)
Rmin x0 (b - a) <= b + - a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x = a
H5:a <= a <= b
H6:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < alp -> Rabs (f0 x2 - f0 a) < eps0)
eps:R
H7:eps > 0
x0:R
H8:x0 > 0 /\ (forall x2 : R, D_x no_cond a x2 /\ Rabs (x2 - a) < x0 -> Rabs (f0 x2 - f0 a) < eps)
x1:R
H9:Rabs (x1 - a) < Rmin x0 (b - a)

Rmin x0 (b - a) <= b + - a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a

continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b

continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b

a <= x <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b

continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps

exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (h x1 - h x) < eps)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps

0 < x - a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (h x1 - h x) < eps)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a

exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (h x1 - h x) < eps)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a

0 < b - x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (h x1 - h x) < eps)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x

exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (h x1 - h x) < eps)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x

Rmin x0 (Rmin (x - a) (b - x)) > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x)) -> Rabs (h x1 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
Hle:x - a <= b - x

(if Rle_dec x0 (x - a) then x0 else x - a) > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
Hnle:~ x - a <= b - x
(if Rle_dec x0 (b - x) then x0 else b - x) > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x)) -> Rabs (h x1 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
Hle:x - a <= b - x
Hlea:x0 <= x - a

x0 > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
Hle:x - a <= b - x
Hnlea:~ x0 <= x - a
x - a > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
Hnle:~ x - a <= b - x
(if Rle_dec x0 (b - x) then x0 else b - x) > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x)) -> Rabs (h x1 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
Hle:x - a <= b - x
Hnlea:~ x0 <= x - a

x - a > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
Hnle:~ x - a <= b - x
(if Rle_dec x0 (b - x) then x0 else b - x) > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x)) -> Rabs (h x1 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
Hnle:~ x - a <= b - x

(if Rle_dec x0 (b - x) then x0 else b - x) > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x)) -> Rabs (h x1 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
Hnle:~ x - a <= b - x
Hleb:x0 <= b - x

x0 > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
Hnle:~ x - a <= b - x
Hnleb:~ x0 <= b - x
b - x > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x)) -> Rabs (h x1 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
Hnle:~ x - a <= b - x
Hnleb:~ x0 <= b - x

b - x > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x)) -> Rabs (h x1 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x

forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x)) -> Rabs (h x1 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))

a < x1 < b -> Rabs (h x1 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
a < x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hle:x <= a

Rabs ((if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b) - f0 a) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
Rabs ((if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b) - (if Rle_dec x b then f0 x else f0 b)) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
a < x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a

Rabs ((if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b) - (if Rle_dec x b then f0 x else f0 b)) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
a < x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
r:x <= b

Rabs ((if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b) - f0 x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
x <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
a < x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
r:x <= b
Hle0:x1 <= a

Rabs (f0 a - f0 x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
r:x <= b
n:~ x1 <= a
Rabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
x <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
a < x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
r:x <= b
n:~ x1 <= a

Rabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
x <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
a < x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
r:x <= b
n:~ x1 <= a
r0:x1 <= b

Rabs (f0 x1 - f0 x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
r:x <= b
n:~ x1 <= a
x1 <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
x <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
a < x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
r:x <= b
n:~ x1 <= a
r0:x1 <= b

D_x no_cond x x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
r:x <= b
n:~ x1 <= a
r0:x1 <= b
Rabs (x1 - x) < x0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
r:x <= b
n:~ x1 <= a
x1 <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
x <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
a < x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
r:x <= b
n:~ x1 <= a
r0:x1 <= b

Rabs (x1 - x) < x0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
r:x <= b
n:~ x1 <= a
x1 <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
x <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
a < x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
r:x <= b
n:~ x1 <= a
r0:x1 <= b

Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
r:x <= b
n:~ x1 <= a
r0:x1 <= b
Rmin x0 (Rmin (x - a) (b - x)) <= x0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
r:x <= b
n:~ x1 <= a
x1 <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
x <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
a < x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
r:x <= b
n:~ x1 <= a
r0:x1 <= b

Rmin x0 (Rmin (x - a) (b - x)) <= x0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
r:x <= b
n:~ x1 <= a
x1 <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
x <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
a < x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
r:x <= b
n:~ x1 <= a

x1 <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a
x <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
a < x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
H15:a < x1
H16:x1 < b
Hnle:~ x <= a

x <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
a < x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))

a < x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))

a < x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))

x + - x1 <= Rabs (x1 - x)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
Rabs (x1 - x) < x + - a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))

Rabs (x1 - x) < x + - a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))

Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
Rmin x0 (Rmin (x - a) (b - x)) <= x + - a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))

Rmin x0 (Rmin (x - a) (b - x)) <= x + - a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))

Rmin x0 (Rmin (x - a) (b - x)) <= Rmin (x - a) (b - x)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
Rmin (x - a) (b - x) <= x + - a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))

Rmin (x - a) (b - x) <= x + - a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))

x1 < b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))

x1 + - x <= Rabs (x1 - x)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
Rabs (x1 - x) < b + - x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))

Rabs (x1 - x) < b + - x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))

Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))
Rmin x0 (Rmin (x - a) (b - x)) <= b + - x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x < b
H6:a <= x <= b
H7:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
eps:R
H8:eps > 0
x0:R
H9:x0 > 0
H10:forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps
H11:0 < x - a
H12:0 < b - x
x1:R
H13:D_x no_cond x x1
H14:Rabs (x1 - x) < Rmin x0 (Rmin (x - a) (b - x))

Rmin x0 (Rmin (x - a) (b - x)) <= b + - x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b

continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b

continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b

a <= b <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b

continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond b x1 /\ Rabs (x1 - b) < alp -> Rabs (f0 x1 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x1 : R, D_x no_cond b x1 /\ Rabs (x1 - b) < x0 -> Rabs (f0 x1 - f0 b) < eps)

Rmin x0 (b - a) > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond b x1 /\ Rabs (x1 - b) < alp -> Rabs (f0 x1 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x1 : R, D_x no_cond b x1 /\ Rabs (x1 - b) < x0 -> Rabs (f0 x1 - f0 b) < eps)
forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (b - a) -> Rabs (h x1 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond b x1 /\ Rabs (x1 - b) < alp -> Rabs (f0 x1 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x1 : R, D_x no_cond b x1 /\ Rabs (x1 - b) < x0 -> Rabs (f0 x1 - f0 b) < eps)
r:x0 <= b - a

x0 > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond b x1 /\ Rabs (x1 - b) < alp -> Rabs (f0 x1 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x1 : R, D_x no_cond b x1 /\ Rabs (x1 - b) < x0 -> Rabs (f0 x1 - f0 b) < eps)
n:~ x0 <= b - a
b - a > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond b x1 /\ Rabs (x1 - b) < alp -> Rabs (f0 x1 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x1 : R, D_x no_cond b x1 /\ Rabs (x1 - b) < x0 -> Rabs (f0 x1 - f0 b) < eps)
forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (b - a) -> Rabs (h x1 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond b x1 /\ Rabs (x1 - b) < alp -> Rabs (f0 x1 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x1 : R, D_x no_cond b x1 /\ Rabs (x1 - b) < x0 -> Rabs (f0 x1 - f0 b) < eps)
n:~ x0 <= b - a

b - a > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond b x1 /\ Rabs (x1 - b) < alp -> Rabs (f0 x1 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x1 : R, D_x no_cond b x1 /\ Rabs (x1 - b) < x0 -> Rabs (f0 x1 - f0 b) < eps)
forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (b - a) -> Rabs (h x1 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond b x1 /\ Rabs (x1 - b) < alp -> Rabs (f0 x1 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x1 : R, D_x no_cond b x1 /\ Rabs (x1 - b) < x0 -> Rabs (f0 x1 - f0 b) < eps)

forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < Rmin x0 (b - a) -> Rabs (h x1 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)

a < x1 -> Rabs (h x1 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
a < x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hlea:x <= a

Rabs ((if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b) - f0 a) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Rabs ((if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b) - (if Rle_dec x b then f0 x else f0 b)) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
a < x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a

Rabs ((if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b) - (if Rle_dec x b then f0 x else f0 b)) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
a < x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hlea':x1 <= a

Rabs (f0 a - (if Rle_dec x b then f0 x else f0 b)) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Rabs ((if Rle_dec x1 b then f0 x1 else f0 b) - (if Rle_dec x b then f0 x else f0 b)) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
a < x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a

Rabs ((if Rle_dec x1 b then f0 x1 else f0 b) - (if Rle_dec x b then f0 x else f0 b)) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
a < x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b

Rabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hnleb:~ x <= b
Rabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
a < x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
Hleb':x1 <= b

Rabs (f0 x1 - f0 x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
Hnleb':~ x1 <= b
Rabs (f0 b - f0 x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hnleb:~ x <= b
Rabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
a < x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
H15:x1 < b
H13:x0 > 0
H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps

Rabs (f0 x1 - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
H15:x1 = b
H13:x0 > 0
H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps
Rabs (f0 x1 - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
Hnleb':~ x1 <= b
Rabs (f0 b - f0 x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hnleb:~ x <= b
Rabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
a < x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
H15:x1 < b
H13:x0 > 0
H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps

D_x no_cond b x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
H15:x1 < b
H13:x0 > 0
H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps
Rabs (x1 - b) < x0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
H15:x1 = b
H13:x0 > 0
H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps
Rabs (f0 x1 - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
Hnleb':~ x1 <= b
Rabs (f0 b - f0 x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hnleb:~ x <= b
Rabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
a < x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
H15:x1 < b
H13:x0 > 0
H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps

True
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
H15:x1 < b
H13:x0 > 0
H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps
b <> x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
H15:x1 < b
H13:x0 > 0
H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps
Rabs (x1 - b) < x0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
H15:x1 = b
H13:x0 > 0
H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps
Rabs (f0 x1 - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
Hnleb':~ x1 <= b
Rabs (f0 b - f0 x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hnleb:~ x <= b
Rabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
a < x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
H15:x1 < b
H13:x0 > 0
H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps

b <> x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
H15:x1 < b
H13:x0 > 0
H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps
Rabs (x1 - b) < x0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
H15:x1 = b
H13:x0 > 0
H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps
Rabs (f0 x1 - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
Hnleb':~ x1 <= b
Rabs (f0 b - f0 x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hnleb:~ x <= b
Rabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
a < x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
H15:x1 < b
H13:x0 > 0
H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps

Rabs (x1 - b) < x0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
H15:x1 = b
H13:x0 > 0
H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps
Rabs (f0 x1 - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
Hnleb':~ x1 <= b
Rabs (f0 b - f0 x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hnleb:~ x <= b
Rabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
a < x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - b) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
H15:x1 < b
H13:x0 > 0
H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps

Rabs (x1 - b) < Rmin x0 (b - a)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - b) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
H15:x1 < b
H13:x0 > 0
H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps
Rmin x0 (b - a) <= x0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
H15:x1 = b
H13:x0 > 0
H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps
Rabs (f0 x1 - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
Hnleb':~ x1 <= b
Rabs (f0 b - f0 x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hnleb:~ x <= b
Rabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
a < x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - b) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
H15:x1 < b
H13:x0 > 0
H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps

Rmin x0 (b - a) <= x0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
H15:x1 = b
H13:x0 > 0
H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps
Rabs (f0 x1 - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
Hnleb':~ x1 <= b
Rabs (f0 b - f0 x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hnleb:~ x <= b
Rabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
a < x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
H15:x1 = b
H13:x0 > 0
H14:forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps

Rabs (f0 x1 - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
Hnleb':~ x1 <= b
Rabs (f0 b - f0 x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hnleb:~ x <= b
Rabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
a < x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hleb:x <= b
Hnleb':~ x1 <= b

Rabs (f0 b - f0 x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hnleb:~ x <= b
Rabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
a < x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
H12:a < x1
Hnlea:~ x <= a
Hnlea':~ x1 <= a
Hnleb:~ x <= b

Rabs ((if Rle_dec x1 b then f0 x1 else f0 b) - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)
a < x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - x) < Rmin x0 (b - a)

a < x1
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - b) < Rmin x0 (b - a)

b + - x1 <= Rabs (x1 - b)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - b) < Rmin x0 (b - a)
Rabs (x1 - b) < b + - a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - b) < Rmin x0 (b - a)

Rabs (x1 - b) < b + - a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - b) < Rmin x0 (b - a)

Rabs (x1 - b) < Rmin x0 (b - a)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - b) < Rmin x0 (b - a)
Rmin x0 (b - a) <= b + - a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x2 : R => if Rle_dec x2 a then f0 a else if Rle_dec x2 b then f0 x2 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x = b
H7:a <= b <= b
H8:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < alp -> Rabs (f0 x2 - f0 b) < eps0)
eps:R
H9:eps > 0
x0:R
H10:x0 > 0 /\ (forall x2 : R, D_x no_cond b x2 /\ Rabs (x2 - b) < x0 -> Rabs (f0 x2 - f0 b) < eps)
x1:R
H11:Rabs (x1 - b) < Rmin x0 (b - a)

Rmin x0 (b - a) <= b + - a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b

continuity_pt h x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0

x - b > 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < x - b -> Rabs (h x0 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x0 : R => if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0

forall x0 : R, D_x no_cond x x0 /\ Rabs (x0 - x) < x - b -> Rabs (h x0 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
x0:R
H8:D_x no_cond x x0
H9:Rabs (x0 - x) < x - b

Rabs (h x0 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
x0:R
H8:D_x no_cond x x0
H9:Rabs (x0 - x) < x - b

b < x0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
x0:R
H8:D_x no_cond x x0
H9:Rabs (x0 - x) < x - b
H10:b < x0
Rabs (h x0 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
x0:R
H8:D_x no_cond x x0
H9:Rabs (x0 - x) < x - b

x + - x0 <= Rabs (x0 - x)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
x0:R
H8:D_x no_cond x x0
H9:Rabs (x0 - x) < x - b
Rabs (x0 - x) < x + - b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
x0:R
H8:D_x no_cond x x0
H9:Rabs (x0 - x) < x - b
H10:b < x0
Rabs (h x0 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
x0:R
H8:D_x no_cond x x0
H9:Rabs (x0 - x) < x - b

Rabs (x0 - x) < x + - b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
x0:R
H8:D_x no_cond x x0
H9:Rabs (x0 - x) < x - b
H10:b < x0
Rabs (h x0 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
x0:R
H8:D_x no_cond x x0
H9:Rabs (x0 - x) < x - b
H10:b < x0

Rabs (h x0 - h x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
x0:R
H8:D_x no_cond x x0
H9:Rabs (x0 - x) < x - b
H10:b < x0
Hle:x <= a

Rabs ((if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b) - f0 a) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
x0:R
H8:D_x no_cond x x0
H9:Rabs (x0 - x) < x - b
H10:b < x0
Hnle:~ x <= a
Rabs ((if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b) - (if Rle_dec x b then f0 x else f0 b)) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
x0:R
H8:D_x no_cond x x0
H9:Rabs (x0 - x) < x - b
H10:b < x0
Hnle:~ x <= a

Rabs ((if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b) - (if Rle_dec x b then f0 x else f0 b)) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
x0:R
H8:D_x no_cond x x0
H9:Rabs (x0 - x) < x - b
H10:b < x0
Hnle:~ x <= a
Hleb:x <= b

Rabs ((if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b) - f0 x) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
x0:R
H8:D_x no_cond x x0
H9:Rabs (x0 - x) < x - b
H10:b < x0
Hnle:~ x <= a
Hnleb:~ x <= b
Rabs ((if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b) - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
x0:R
H8:D_x no_cond x x0
H9:Rabs (x0 - x) < x - b
H10:b < x0
Hnle:~ x <= a
Hnleb:~ x <= b

Rabs ((if Rle_dec x0 a then f0 a else if Rle_dec x0 b then f0 x0 else f0 b) - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
x0:R
H8:D_x no_cond x x0
H9:Rabs (x0 - x) < x - b
H10:b < x0
Hnle:~ x <= a
Hnleb:~ x <= b
Hlea':x0 <= a

Rabs (f0 a - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
x0:R
H8:D_x no_cond x x0
H9:Rabs (x0 - x) < x - b
H10:b < x0
Hnle:~ x <= a
Hnleb:~ x <= b
Hnlea':~ x0 <= a
Rabs ((if Rle_dec x0 b then f0 x0 else f0 b) - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
x0:R
H8:D_x no_cond x x0
H9:Rabs (x0 - x) < x - b
H10:b < x0
Hnle:~ x <= a
Hnleb:~ x <= b
Hnlea':~ x0 <= a

Rabs ((if Rle_dec x0 b then f0 x0 else f0 b) - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
x0:R
H8:D_x no_cond x x0
H9:Rabs (x0 - x) < x - b
H10:b < x0
Hnle:~ x <= a
Hnleb:~ x <= b
Hnlea':~ x0 <= a
Hleb':x0 <= b

Rabs (f0 x0 - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
x0:R
H8:D_x no_cond x x0
H9:Rabs (x0 - x) < x - b
H10:b < x0
Hnle:~ x <= a
Hnleb:~ x <= b
Hnlea':~ x0 <= a
Hnleb':~ x0 <= b
Rabs (f0 b - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x1 : R => if Rle_dec x1 a then f0 a else if Rle_dec x1 b then f0 x1 else f0 b:R -> R
H2:0 < b - a
x:R
H3:x = a \/ x > a
H4:x > a
H5:x = b \/ x > b
H6:x > b
eps:R
H7:eps > 0
x0:R
H8:D_x no_cond x x0
H9:Rabs (x0 - x) < x - b
H10:b < x0
Hnle:~ x <= a
Hnleb:~ x <= b
Hnlea':~ x0 <= a
Hnleb':~ x0 <= b

Rabs (f0 b - f0 b) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a

forall c : R, a <= c <= b -> h c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
c:R
H3:a <= c <= b
H4:a <= c
H5:c <= b
H6:c < a

f0 a = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
c:R
H3:a <= c <= b
H4:a <= c
H5:c <= b
H6:c = a
f0 a = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
c:R
H3:a <= c <= b
H4:a <= c
H5:c <= b
n:~ c <= a
(if Rle_dec c b then f0 c else f0 b) = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
c:R
H3:a <= c <= b
H4:a <= c
H5:c <= b
H6:c = a

f0 a = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
c:R
H3:a <= c <= b
H4:a <= c
H5:c <= b
n:~ c <= a
(if Rle_dec c b then f0 c else f0 b) = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
c:R
H3:a <= c <= b
H4:a <= c
H5:c <= b
n:~ c <= a

(if Rle_dec c b then f0 c else f0 b) = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
c:R
H3:a <= c <= b
H4:a <= c
H5:c <= b
n:~ c <= a
r:c <= b

f0 c = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
c:R
H3:a <= c <= b
H4:a <= c
H5:c <= b
n:~ c <= a
c <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0
H1:a < b
h:=fun x : R => if Rle_dec x a then f0 a else if Rle_dec x b then f0 x else f0 b:R -> R
H2:0 < b - a
c:R
H3:a <= c <= b
H4:a <= c
H5:c <= b
n:~ c <= a

c <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b

exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b

continuity (fun _ : R => f0 a)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b
forall c : R, a <= c <= b -> f0 a = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:a = b

forall c : R, a <= c <= b -> f0 a = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0
H1:a = b
c:R
H2:a <= c <= b
H3:b <= c
H4:c <= b

b = c -> f0 a = f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0
H1:a = b
c:R
H2:a <= c <= b
H3:b <= c
H4:c <= b
b = c
f0:R -> R
a, b:R
H:a <= b
H0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0
H1:a = b
c:R
H2:a <= c <= b
H3:b <= c
H4:c <= b

b = c
apply Rle_antisym; assumption. Qed. (**********)

forall (f0 : R -> R) (a b : R), a <= b -> (forall c : R, a <= c <= b -> continuity_pt f0 c) -> exists Mx : R, (forall c : R, a <= c <= b -> f0 c <= f0 Mx) /\ a <= Mx <= b

forall (f0 : R -> R) (a b : R), a <= b -> (forall c : R, a <= c <= b -> continuity_pt f0 c) -> exists Mx : R, (forall c : R, a <= c <= b -> f0 c <= f0 Mx) /\ a <= Mx <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c

(exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)) -> exists Mx : R, (forall c : R, a <= c <= b -> f0 c <= f0 Mx) /\ a <= Mx <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)

exists Mx : R, (forall c : R, a <= c <= b -> f0 c <= f0 Mx) /\ a <= Mx <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont_eq:continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)

exists Mx : R, (forall c : R, a <= c <= b -> f0 c <= f0 Mx) /\ a <= Mx <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c

exists Mx : R, (forall c : R, a <= c <= b -> f0 c <= f0 Mx) /\ a <= Mx <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)

exists Mx : R, (forall c : R, a <= c <= b -> f0 c <= f0 Mx) /\ a <= Mx <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))

exists Mx : R, (forall c : R, a <= c <= b -> f0 c <= f0 Mx) /\ a <= Mx <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))

exists Mx : R, (forall c : R, a <= c <= b -> f0 c <= f0 Mx) /\ a <= Mx <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))

exists Mx : R, (forall c : R, a <= c <= b -> f0 c <= f0 Mx) /\ a <= Mx <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))

bound (image_dir g (fun c : R => a <= c <= b)) -> exists Mx : R, (forall c : R, a <= c <= b -> f0 c <= f0 Mx) /\ a <= Mx <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))

(exists x : R, image_dir g (fun c : R => a <= c <= b) x) -> bound (image_dir g (fun c : R => a <= c <= b)) -> exists Mx : R, (forall c : R, a <= c <= b -> f0 c <= f0 Mx) /\ a <= Mx <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
H7:{m : R | is_lub (image_dir g (fun c : R => a <= c <= b)) m}

exists Mx : R, (forall c : R, a <= c <= b -> f0 c <= f0 Mx) /\ a <= Mx <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M

image_dir g (fun c : R => a <= c <= b) M -> exists Mx : R, (forall c : R, a <= c <= b -> f0 c <= f0 Mx) /\ a <= Mx <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
image_dir g (fun c : R => a <= c <= b) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
Mxx:R
H8:M = g Mxx
H9:a <= Mxx <= b

forall c : R, a <= c <= b -> f0 c <= f0 Mxx
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
Mxx:R
H8:M = g Mxx
H9:a <= Mxx <= b
a <= Mxx <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
image_dir g (fun c : R => a <= c <= b) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
Mxx:R
H8:M = g Mxx
H9:a <= Mxx <= b

a <= Mxx <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
image_dir g (fun c : R => a <= c <= b) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M

image_dir g (fun c : R => a <= c <= b) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:image_dir g (fun c : R => a <= c <= b) M

image_dir g (fun c : R => a <= c <= b) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
image_dir g (fun c : R => a <= c <= b) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M

image_dir g (fun c : R => a <= c <= b) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M

(exists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y) -> image_dir g (fun c : R => a <= c <= b) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0

is_upper_bound (image_dir g (fun c : R => a <= c <= b)) (M - eps) -> image_dir g (fun c : R => a <= c <= b) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
is_upper_bound (image_dir g (fun c : R => a <= c <= b)) (M - eps)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
H11:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) (M - eps)
H12:M <= M - eps

M - eps < M -> image_dir g (fun c : R => a <= c <= b) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
H11:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) (M - eps)
H12:M <= M - eps
M - eps < M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
is_upper_bound (image_dir g (fun c : R => a <= c <= b)) (M - eps)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
H11:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) (M - eps)
H12:M <= M - eps

M - eps < M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
is_upper_bound (image_dir g (fun c : R => a <= c <= b)) (M - eps)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0

is_upper_bound (image_dir g (fun c : R => a <= c <= b)) (M - eps)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b

x <= M -> x <= M - eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
x <= M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
H13:x <= M - eps

x <= M - eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps
x <= M - eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
x <= M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps

x <= M - eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
x <= M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps

Rabs (x - M) < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps
exists y : R, x = g y /\ a <= y <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
x <= M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps

M - x < eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps
M - x >= 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps
exists y : R, x = g y /\ a <= y <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
x <= M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps

M - eps < x - eps + eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps
M - eps = x - eps + (M - x)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps
M - x >= 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps
exists y : R, x = g y /\ a <= y <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
x <= M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps

M - eps < x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps
x = x - eps + eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps
M - eps = x - eps + (M - x)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps
M - x >= 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps
exists y : R, x = g y /\ a <= y <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
x <= M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps

x = x - eps + eps
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps
M - eps = x - eps + (M - x)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps
M - x >= 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps
exists y : R, x = g y /\ a <= y <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
x <= M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps

M - eps = x - eps + (M - x)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps
M - x >= 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps
exists y : R, x = g y /\ a <= y <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
x <= M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps

M - x >= 0
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps
exists y : R, x = g y /\ a <= y <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
x <= M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
H12:x <= M
n:~ x <= M - eps

exists y : R, x = g y /\ a <= y <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b
x <= M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x0 : R, image_dir g (fun c : R => a <= c <= b) x0
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H8:~ image_dir g (fun c : R => a <= c <= b) M
eps:posreal
H9:forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
H7:is_upper_bound (image_dir g (fun c : R => a <= c <= b)) M
H10:forall b0 : R, is_upper_bound (image_dir g (fun c : R => a <= c <= b)) b0 -> M <= b0
x:R
H11:exists y : R, x = g y /\ a <= y <= b

x <= M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M

exists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M

(exists V : R -> Prop, neighbourhood V M /\ (forall y : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)) -> exists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists V : R -> Prop, neighbourhood V M /\ (forall y : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
H9:exists V0 : R -> Prop, neighbourhood V0 M /\ (forall y : R, ~ intersection_domain V0 (image_dir g (fun c : R => a <= c <= b)) y)
V:R -> Prop
H10:neighbourhood V M
H11:forall y : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y

exists eps : posreal, forall y : R, ~ intersection_domain (disc M eps) (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists V : R -> Prop, neighbourhood V M /\ (forall y : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
H9:exists V0 : R -> Prop, neighbourhood V0 M /\ (forall y0 : R, ~ intersection_domain V0 (image_dir g (fun c : R => a <= c <= b)) y0)
V:R -> Prop
H10:exists delta : posreal, included (disc M delta) V
H11:forall y0 : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y0
del:posreal
H12:included (disc M del) V
y:R
H13:intersection_domain (disc M del) (image_dir g (fun c : R => a <= c <= b)) y

intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists V : R -> Prop, neighbourhood V M /\ (forall y : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
H9:exists V0 : R -> Prop, neighbourhood V0 M /\ (forall y0 : R, ~ intersection_domain V0 (image_dir g (fun c : R => a <= c <= b)) y0)
V:R -> Prop
H10:exists delta : posreal, included (disc M delta) V
H11:forall y0 : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y0
del:posreal
H12:included (disc M del) V
y:R
H13:disc M del y
H14:image_dir g (fun c : R => a <= c <= b) y

V y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
H9:exists V0 : R -> Prop, neighbourhood V0 M /\ (forall y0 : R, ~ intersection_domain V0 (image_dir g (fun c : R => a <= c <= b)) y0)
V:R -> Prop
H10:exists delta : posreal, included (disc M delta) V
H11:forall y0 : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y0
del:posreal
H12:included (disc M del) V
y:R
H13:disc M del y
H14:image_dir g (fun c : R => a <= c <= b) y
image_dir g (fun c : R => a <= c <= b) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists V : R -> Prop, neighbourhood V M /\ (forall y : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
H9:exists V0 : R -> Prop, neighbourhood V0 M /\ (forall y0 : R, ~ intersection_domain V0 (image_dir g (fun c : R => a <= c <= b)) y0)
V:R -> Prop
H10:exists delta : posreal, included (disc M delta) V
H11:forall y0 : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y0
del:posreal
H12:included (disc M del) V
y:R
H13:disc M del y
H14:image_dir g (fun c : R => a <= c <= b) y

image_dir g (fun c : R => a <= c <= b) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
exists V : R -> Prop, neighbourhood V M /\ (forall y : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M

exists V : R -> Prop, neighbourhood V M /\ (forall y : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M

~ point_adherent (image_dir g (fun c : R => a <= c <= b)) M -> exists V : R -> Prop, neighbourhood V M /\ (forall y : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
~ point_adherent (image_dir g (fun c : R => a <= c <= b)) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
H9:~ (forall V : R -> Prop, neighbourhood V M -> exists y : R, intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)

exists V : R -> Prop, neighbourhood V M /\ (forall y : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
~ point_adherent (image_dir g (fun c : R => a <= c <= b)) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
H9:~ (forall V : R -> Prop, neighbourhood V M -> exists y : R, intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)
H10:exists n : R -> Prop, ~ (fun V : R -> Prop => neighbourhood V M -> exists y : R, intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y) n

exists V : R -> Prop, neighbourhood V M /\ (forall y : R, ~ intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
~ point_adherent (image_dir g (fun c : R => a <= c <= b)) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
H9:~ (forall V : R -> Prop, neighbourhood V M -> exists y : R, intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)
H10:exists n : R -> Prop, ~ (fun V : R -> Prop => neighbourhood V M -> exists y : R, intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y) n
V0:R -> Prop
H11:~ (neighbourhood V0 M -> exists y : R, intersection_domain V0 (image_dir g (fun c : R => a <= c <= b)) y)
H12:neighbourhood V0 M
H13:~ (exists y : R, intersection_domain V0 (image_dir g (fun c : R => a <= c <= b)) y)

neighbourhood V0 M /\ (forall y : R, ~ intersection_domain V0 (image_dir g (fun c : R => a <= c <= b)) y)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
~ point_adherent (image_dir g (fun c : R => a <= c <= b)) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
H9:~ (forall V : R -> Prop, neighbourhood V M -> exists y : R, intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)
H10:exists n : R -> Prop, ~ (fun V : R -> Prop => neighbourhood V M -> exists y : R, intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y) n
V0:R -> Prop
H11:~ (neighbourhood V0 M -> exists y : R, intersection_domain V0 (image_dir g (fun c : R => a <= c <= b)) y)
H12:neighbourhood V0 M
H13:~ (exists y : R, intersection_domain V0 (image_dir g (fun c : R => a <= c <= b)) y)

neighbourhood V0 M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
H9:~ (forall V : R -> Prop, neighbourhood V M -> exists y : R, intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)
H10:exists n : R -> Prop, ~ (fun V : R -> Prop => neighbourhood V M -> exists y : R, intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y) n
V0:R -> Prop
H11:~ (neighbourhood V0 M -> exists y : R, intersection_domain V0 (image_dir g (fun c : R => a <= c <= b)) y)
H12:neighbourhood V0 M
H13:~ (exists y : R, intersection_domain V0 (image_dir g (fun c : R => a <= c <= b)) y)
forall y : R, ~ intersection_domain V0 (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
~ point_adherent (image_dir g (fun c : R => a <= c <= b)) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
H9:~ (forall V : R -> Prop, neighbourhood V M -> exists y : R, intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y)
H10:exists n : R -> Prop, ~ (fun V : R -> Prop => neighbourhood V M -> exists y : R, intersection_domain V (image_dir g (fun c : R => a <= c <= b)) y) n
V0:R -> Prop
H11:~ (neighbourhood V0 M -> exists y : R, intersection_domain V0 (image_dir g (fun c : R => a <= c <= b)) y)
H12:neighbourhood V0 M
H13:~ (exists y : R, intersection_domain V0 (image_dir g (fun c : R => a <= c <= b)) y)

forall y : R, ~ intersection_domain V0 (image_dir g (fun c : R => a <= c <= b)) y
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
~ point_adherent (image_dir g (fun c : R => a <= c <= b)) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M

~ point_adherent (image_dir g (fun c : R => a <= c <= b)) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
H9:point_adherent (image_dir g (fun c : R => a <= c <= b)) M

adherence (image_dir g (fun c : R => a <= c <= b)) M -> False
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
H9:point_adherent (image_dir g (fun c : R => a <= c <= b)) M
adherence (image_dir g (fun c : R => a <= c <= b)) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
H9:point_adherent (image_dir g (fun c : R => a <= c <= b)) M
H10:adherence (image_dir g (fun c : R => a <= c <= b)) M
H11:closed_set (image_dir g (fun c : R => a <= c <= b)) -> image_dir g (fun c : R => a <= c <= b) =_D adherence (image_dir g (fun c : R => a <= c <= b))
H12:image_dir g (fun c : R => a <= c <= b) =_D adherence (image_dir g (fun c : R => a <= c <= b))

False
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
H9:point_adherent (image_dir g (fun c : R => a <= c <= b)) M
adherence (image_dir g (fun c : R => a <= c <= b)) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
H9:point_adherent (image_dir g (fun c : R => a <= c <= b)) M
H10:adherence (image_dir g (fun c : R => a <= c <= b)) M
H11:closed_set (image_dir g (fun c : R => a <= c <= b)) -> image_dir g (fun c : R => a <= c <= b) =_D adherence (image_dir g (fun c : R => a <= c <= b))
H12:image_dir g (fun c : R => a <= c <= b) =_D adherence (image_dir g (fun c : R => a <= c <= b))

image_dir g (fun c : R => a <= c <= b) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
H9:point_adherent (image_dir g (fun c : R => a <= c <= b)) M
adherence (image_dir g (fun c : R => a <= c <= b)) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
H9:point_adherent (image_dir g (fun c : R => a <= c <= b)) M
H10:adherence (image_dir g (fun c : R => a <= c <= b)) M
H11:closed_set (image_dir g (fun c : R => a <= c <= b)) -> image_dir g (fun c : R => a <= c <= b) =_D adherence (image_dir g (fun c : R => a <= c <= b))
H12:included (image_dir g (fun c : R => a <= c <= b)) (adherence (image_dir g (fun c : R => a <= c <= b)))
H13:included (adherence (image_dir g (fun c : R => a <= c <= b))) (image_dir g (fun c : R => a <= c <= b))

image_dir g (fun c : R => a <= c <= b) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
H9:point_adherent (image_dir g (fun c : R => a <= c <= b)) M
adherence (image_dir g (fun c : R => a <= c <= b)) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
H5:exists x : R, image_dir g (fun c : R => a <= c <= b) x
H6:bound (image_dir g (fun c : R => a <= c <= b))
M:R
H7:is_lub (image_dir g (fun c : R => a <= c <= b)) M
H8:~ image_dir g (fun c : R => a <= c <= b) M
H9:point_adherent (image_dir g (fun c : R => a <= c <= b)) M

adherence (image_dir g (fun c : R => a <= c <= b)) M
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))

exists x : R, image_dir g (fun c : R => a <= c <= b) x
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))

g a = g a
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
a <= a <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))

a <= a <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))
bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
HypProl:exists g0 : R -> R, continuity g0 /\ (forall c : R, a <= c <= b -> g0 c = f0 c)
g:R -> R
Hcont:continuity g
Heq:forall c : R, a <= c <= b -> g c = f0 c
H1:compact (fun c : R => a <= c <= b)
H2:compact (image_dir g (fun c : R => a <= c <= b))
H3:closed_set (image_dir g (fun c : R => a <= c <= b))
H4:bounded (image_dir g (fun c : R => a <= c <= b))

bound (image_dir g (fun c : R => a <= c <= b))
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c

exists g : R -> R, continuity g /\ (forall c : R, a <= c <= b -> g c = f0 c)
apply prolongement_C0; assumption. Qed. (**********)

forall (f0 : R -> R) (a b : R), a <= b -> (forall c : R, a <= c <= b -> continuity_pt f0 c) -> exists mx : R, (forall c : R, a <= c <= b -> f0 mx <= f0 c) /\ a <= mx <= b

forall (f0 : R -> R) (a b : R), a <= b -> (forall c : R, a <= c <= b -> continuity_pt f0 c) -> exists mx : R, (forall c : R, a <= c <= b -> f0 mx <= f0 c) /\ a <= mx <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c

exists mx : R, (forall c : R, a <= c <= b -> f0 mx <= f0 c) /\ a <= mx <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c

(forall c : R, a <= c <= b -> continuity_pt (- f0) c) -> exists mx : R, (forall c : R, a <= c <= b -> f0 mx <= f0 c) /\ a <= mx <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
forall c : R, a <= c <= b -> continuity_pt (- f0) c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:forall c : R, a <= c <= b -> continuity_pt (- f0) c
H2:exists Mx : R, (forall c : R, a <= c <= b -> (- f0)%F c <= (- f0)%F Mx) /\ a <= Mx <= b
x0:R
H3:(forall c : R, a <= c <= b -> (- f0)%F c <= (- f0)%F x0) /\ a <= x0 <= b

forall c : R, a <= c <= b -> f0 x0 <= f0 c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:forall c : R, a <= c <= b -> continuity_pt (- f0) c
H2:exists Mx : R, (forall c : R, a <= c <= b -> (- f0)%F c <= (- f0)%F Mx) /\ a <= Mx <= b
x0:R
H3:(forall c : R, a <= c <= b -> (- f0)%F c <= (- f0)%F x0) /\ a <= x0 <= b
a <= x0 <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
forall c : R, a <= c <= b -> continuity_pt (- f0) c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
H1:forall c : R, a <= c <= b -> continuity_pt (- f0) c
H2:exists Mx : R, (forall c : R, a <= c <= b -> (- f0)%F c <= (- f0)%F Mx) /\ a <= Mx <= b
x0:R
H3:(forall c : R, a <= c <= b -> (- f0)%F c <= (- f0)%F x0) /\ a <= x0 <= b

a <= x0 <= b
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c
forall c : R, a <= c <= b -> continuity_pt (- f0) c
f0:R -> R
a, b:R
H:a <= b
H0:forall c : R, a <= c <= b -> continuity_pt f0 c

forall c : R, a <= c <= b -> continuity_pt (- f0) c
f0:R -> R
a, b:R
H:a <= b
H0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0
c:R
H1:a <= c <= b

continuity_pt (- f0) c
f0:R -> R
a, b:R
H:a <= b
H0:forall c0 : R, a <= c0 <= b -> continuity_pt f0 c0
c:R
H1:a <= c <= b
H2:continuity_pt f0 c

continuity_pt (- f0) c
apply (continuity_pt_opp _ _ H2). Qed. (********************************************************)

Proof of Bolzano-Weierstrass theorem

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

Definition ValAdh (un:nat -> R) (x:R) : Prop :=
  forall (V:R -> Prop) (N:nat),
    neighbourhood V x ->  exists p : nat, (N <= p)%nat /\ V (un p).

Definition intersection_family (f:family) (x:R) : Prop :=
  forall y:R, ind f y -> f y x.


forall un : nat -> R, let D := fun x : R => exists n : nat, x = INR n in let f0 := fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x) in forall x : R, (exists y : R, f0 x y) -> D x

forall un : nat -> R, let D := fun x : R => exists n : nat, x = INR n in let f0 := fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x) in forall x : R, (exists y : R, f0 x y) -> D x
un:nat -> R
D:=fun x1 : R => exists n : nat, x1 = INR n:R -> Prop
f0:=fun x1 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x1 <= INR p) /\ D x1):R -> R -> Prop
x:R
H:exists y : R, f0 x y
x0:R
H0:f0 x x0

neighbourhood (disc x0 {| pos := 1; cond_pos := Rlt_0_1 |}) x0
un:nat -> R
D:=fun x1 : R => exists n : nat, x1 = INR n:R -> Prop
f0:=fun x1 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x1 <= INR p) /\ D x1):R -> R -> Prop
x:R
H:exists y : R, f0 x y
x0:R
H0:f0 x x0
H1:neighbourhood (disc x0 {| pos := 1; cond_pos := Rlt_0_1 |}) x0
D x
un:nat -> R
D:=fun x1 : R => exists n : nat, x1 = INR n:R -> Prop
f0:=fun x1 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x1 <= INR p) /\ D x1):R -> R -> Prop
x:R
H:exists y : R, f0 x y
x0:R
H0:f0 x x0
H1:neighbourhood (disc x0 {| pos := 1; cond_pos := Rlt_0_1 |}) x0

D x
elim (H0 _ H1); intros; unfold intersection_domain in H2; elim H2; intros; elim H4; intros; apply H6. Qed. Definition ValAdh_un (un:nat -> R) : R -> Prop := let D := fun x:R => exists n : nat, x = INR n in let f := fun x:R => adherence (fun y:R => (exists p : nat, y = un p /\ x <= INR p) /\ D x) in intersection_family (mkfamily D f (ValAdh_un_exists un)).

forall (un : nat -> R) (x : R), ValAdh un x <-> ValAdh_un un x

forall (un : nat -> R) (x : R), ValAdh un x <-> ValAdh_un un x
un:nat -> R
x:R
H:ValAdh un x

ValAdh_un un x
un:nat -> R
x:R
H:ValAdh_un un x
ValAdh un x
un:nat -> R
x:R
H:forall (V0 : R -> Prop) (N0 : nat), neighbourhood V0 x -> exists p : nat, (N0 <= p)%nat /\ V0 (un p)
y:R
H0:exists n : nat, y = INR n
N:nat
H1:y = INR N
V:R -> Prop
H2:neighbourhood V x
x0:nat
H3:(N <= x0)%nat
H4:V (un x0)

V (un x0)
un:nat -> R
x:R
H:forall (V0 : R -> Prop) (N0 : nat), neighbourhood V0 x -> exists p : nat, (N0 <= p)%nat /\ V0 (un p)
y:R
H0:exists n : nat, y = INR n
N:nat
H1:y = INR N
V:R -> Prop
H2:neighbourhood V x
x0:nat
H3:(N <= x0)%nat
H4:V (un x0)
(exists p : nat, un x0 = un p /\ y <= INR p) /\ (exists n : nat, y = INR n)
un:nat -> R
x:R
H:ValAdh_un un x
ValAdh un x
un:nat -> R
x:R
H:forall (V0 : R -> Prop) (N0 : nat), neighbourhood V0 x -> exists p : nat, (N0 <= p)%nat /\ V0 (un p)
y:R
H0:exists n : nat, y = INR n
N:nat
H1:y = INR N
V:R -> Prop
H2:neighbourhood V x
x0:nat
H3:(N <= x0)%nat
H4:V (un x0)

(exists p : nat, un x0 = un p /\ y <= INR p) /\ (exists n : nat, y = INR n)
un:nat -> R
x:R
H:ValAdh_un un x
ValAdh un x
un:nat -> R
x:R
H:forall (V0 : R -> Prop) (N0 : nat), neighbourhood V0 x -> exists p : nat, (N0 <= p)%nat /\ V0 (un p)
y:R
H0:exists n : nat, y = INR n
N:nat
H1:y = INR N
V:R -> Prop
H2:neighbourhood V x
x0:nat
H3:(N <= x0)%nat
H4:V (un x0)

exists p : nat, un x0 = un p /\ y <= INR p
un:nat -> R
x:R
H:forall (V0 : R -> Prop) (N0 : nat), neighbourhood V0 x -> exists p : nat, (N0 <= p)%nat /\ V0 (un p)
y:R
H0:exists n : nat, y = INR n
N:nat
H1:y = INR N
V:R -> Prop
H2:neighbourhood V x
x0:nat
H3:(N <= x0)%nat
H4:V (un x0)
exists n : nat, y = INR n
un:nat -> R
x:R
H:ValAdh_un un x
ValAdh un x
un:nat -> R
x:R
H:forall (V0 : R -> Prop) (N0 : nat), neighbourhood V0 x -> exists p : nat, (N0 <= p)%nat /\ V0 (un p)
y:R
H0:exists n : nat, y = INR n
N:nat
H1:y = INR N
V:R -> Prop
H2:neighbourhood V x
x0:nat
H3:(N <= x0)%nat
H4:V (un x0)

exists n : nat, y = INR n
un:nat -> R
x:R
H:ValAdh_un un x
ValAdh un x
un:nat -> R
x:R
H:ValAdh_un un x

ValAdh un x
un:nat -> R
x:R
H:forall y : R, (exists n : nat, y = INR n) -> adherence (fun y0 : R => (exists p : nat, y0 = un p /\ y <= INR p) /\ (exists n : nat, y = INR n)) x
V:R -> Prop
N:nat
H0:neighbourhood V x

adherence (fun y0 : R => (exists p : nat, y0 = un p /\ INR N <= INR p) /\ (exists n : nat, INR N = INR n)) x
un:nat -> R
x:R
H:forall y : R, (exists n : nat, y = INR n) -> adherence (fun y0 : R => (exists p : nat, y0 = un p /\ y <= INR p) /\ (exists n : nat, y = INR n)) x
V:R -> Prop
N:nat
H0:neighbourhood V x
H1:adherence (fun y0 : R => (exists p : nat, y0 = un p /\ INR N <= INR p) /\ (exists n : nat, INR N = INR n)) x
exists p : nat, (N <= p)%nat /\ V (un p)
un:nat -> R
x:R
H:forall y : R, (exists n : nat, y = INR n) -> adherence (fun y0 : R => (exists p : nat, y0 = un p /\ y <= INR p) /\ (exists n : nat, y = INR n)) x
V:R -> Prop
N:nat
H0:neighbourhood V x
H1:adherence (fun y0 : R => (exists p : nat, y0 = un p /\ INR N <= INR p) /\ (exists n : nat, INR N = INR n)) x

exists p : nat, (N <= p)%nat /\ V (un p)
un:nat -> R
x:R
H:forall y : R, (exists n : nat, y = INR n) -> adherence (fun y0 : R => (exists p : nat, y0 = un p /\ y <= INR p) /\ (exists n : nat, y = INR n)) x
V:R -> Prop
N:nat
H0:neighbourhood V x
H1:forall V0 : R -> Prop, neighbourhood V0 x -> exists y : R, intersection_domain V0 (fun y0 : R => (exists p : nat, y0 = un p /\ INR N <= INR p) /\ (exists n : nat, INR N = INR n)) y
H2:exists y : R, intersection_domain V (fun y0 : R => (exists p : nat, y0 = un p /\ INR N <= INR p) /\ (exists n : nat, INR N = INR n)) y
x0:R
H3:V x0
H5:exists n : nat, INR N = INR n
x1:nat
H4:x0 = un x1
H6:INR N <= INR x1

(N <= x1)%nat
un:nat -> R
x:R
H:forall y : R, (exists n : nat, y = INR n) -> adherence (fun y0 : R => (exists p : nat, y0 = un p /\ y <= INR p) /\ (exists n : nat, y = INR n)) x
V:R -> Prop
N:nat
H0:neighbourhood V x
H1:forall V0 : R -> Prop, neighbourhood V0 x -> exists y : R, intersection_domain V0 (fun y0 : R => (exists p : nat, y0 = un p /\ INR N <= INR p) /\ (exists n : nat, INR N = INR n)) y
H2:exists y : R, intersection_domain V (fun y0 : R => (exists p : nat, y0 = un p /\ INR N <= INR p) /\ (exists n : nat, INR N = INR n)) y
x0:R
H3:V x0
H5:exists n : nat, INR N = INR n
x1:nat
H4:x0 = un x1
H6:INR N <= INR x1
V (un x1)
un:nat -> R
x:R
H:forall y : R, (exists n : nat, y = INR n) -> adherence (fun y0 : R => (exists p : nat, y0 = un p /\ y <= INR p) /\ (exists n : nat, y = INR n)) x
V:R -> Prop
N:nat
H0:neighbourhood V x
H1:forall V0 : R -> Prop, neighbourhood V0 x -> exists y : R, intersection_domain V0 (fun y0 : R => (exists p : nat, y0 = un p /\ INR N <= INR p) /\ (exists n : nat, INR N = INR n)) y
H2:exists y : R, intersection_domain V (fun y0 : R => (exists p : nat, y0 = un p /\ INR N <= INR p) /\ (exists n : nat, INR N = INR n)) y
x0:R
H3:V x0
H5:exists n : nat, INR N = INR n
x1:nat
H4:x0 = un x1
H6:INR N <= INR x1

V (un x1)
rewrite H4 in H3; apply H3. Qed.

forall F G : R -> Prop, included F G -> included (adherence F) (adherence G)

forall F G : R -> Prop, included F G -> included (adherence F) (adherence G)
unfold adherence, included; unfold point_adherent; intros; elim (H0 _ H1); unfold intersection_domain; intros; elim H2; clear H2; intros; exists x0; split; [ assumption | apply (H _ H3) ]. Qed. Definition family_closed_set (f:family) : Prop := forall x:R, closed_set (f x). Definition intersection_vide_in (D:R -> Prop) (f:family) : Prop := forall x:R, (ind f x -> included (f x) D) /\ ~ (exists y : R, intersection_family f y). Definition intersection_vide_finite_in (D:R -> Prop) (f:family) : Prop := intersection_vide_in D f /\ family_finite f. (**********)

forall X : R -> Prop, compact X -> (exists z : R, X z) -> forall g : family, family_closed_set g -> intersection_vide_in X g -> exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)

forall X : R -> Prop, compact X -> (exists z : R, X z) -> forall g : family, family_closed_set g -> intersection_vide_in X g -> exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)
X:R -> Prop
H:compact X
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g

exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)
X:R -> Prop
H:compact X
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop

exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)
X:R -> Prop
H:compact X
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop

exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)
X:R -> Prop
H:compact X
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop

forall x : R, (exists y : R, f' x y) -> D' x
X:R -> Prop
H:compact X
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)
X:R -> Prop
H:compact X
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x

exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)
X:R -> Prop
H:compact X
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family

exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family

covering_open_set X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family

covering X f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
family_open_set f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family

family_open_set f0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
x:R
H3:D' x

open_set (f0 x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
x:R
H3:~ D' x
open_set (f0 x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
x:R
H3:D' x

open_set (complementary (g x))
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
x:R
H3:D' x
complementary (g x) =_D f0 x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
x:R
H3:~ D' x
open_set (f0 x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
x:R
H3:D' x

complementary (g x) =_D f0 x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
x:R
H3:~ D' x
open_set (f0 x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
x:R
H3:D' x

included (complementary (g x)) (fun y : R => complementary (g x) y /\ D' x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
x:R
H3:D' x
included (fun y : R => complementary (g x) y /\ D' x) (complementary (g x))
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
x:R
H3:~ D' x
open_set (f0 x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
x:R
H3:D' x

included (fun y : R => complementary (g x) y /\ D' x) (complementary (g x))
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
x:R
H3:~ D' x
open_set (f0 x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
x:R
H3:~ D' x

open_set (f0 x)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
x:R
H3:~ D' x

open_set (fun _ : R => False)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
x:R
H3:~ D' x
(fun _ : R => False) =_D f0 x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
x:R
H3:~ D' x

(fun _ : R => False) =_D f0 x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0

exists D : R -> Prop, intersection_vide_finite_in X (subfamily g D)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)

intersection_vide_in X (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R

intersection_domain (ind g) SF x -> included (fun y : R => g x y /\ SF x) X
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
~ (exists y : R, intersection_family (subfamily g SF) y)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:forall x1 : R, (ind g x1 -> included (g x1) X) /\ ~ (exists y : R, intersection_family g y)
D':=ind g:R -> Prop
f':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> Prop
H2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
H5:intersection_domain (ind g) SF x
x0:R
H6:g x x0 /\ SF x
H7:ind g x -> included (g x) X
H8:~ (exists y : R, intersection_family g y)
H9:g x x0
H10:SF x

ind g x
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:forall x1 : R, (ind g x1 -> included (g x1) X) /\ ~ (exists y : R, intersection_family g y)
D':=ind g:R -> Prop
f':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> Prop
H2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
H5:intersection_domain (ind g) SF x
x0:R
H6:g x x0 /\ SF x
H7:ind g x -> included (g x) X
H8:~ (exists y : R, intersection_family g y)
H9:g x x0
H10:SF x
g x x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
~ (exists y : R, intersection_family (subfamily g SF) y)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:forall x1 : R, (ind g x1 -> included (g x1) X) /\ ~ (exists y : R, intersection_family g y)
D':=ind g:R -> Prop
f':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> Prop
H2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
H5:intersection_domain (ind g) SF x
x0:R
H6:g x x0 /\ SF x
H7:ind g x -> included (g x) X
H8:~ (exists y : R, intersection_family g y)
H9:g x x0
H10:SF x

g x x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
~ (exists y : R, intersection_family (subfamily g SF) y)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R

~ (exists y : R, intersection_family (subfamily g SF) y)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y

~ (exists y : R, intersection_family (subfamily g SF) y)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)
~ (exists y : R, intersection_family (subfamily g SF) y)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> Prop
H2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF y

False
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)
~ (exists y : R, intersection_family (subfamily g SF) y)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> Prop
H2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF y

X x0 -> False
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> Prop
H2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF y
X x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)
~ (exists y : R, intersection_family (subfamily g SF) y)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> Prop
H2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, ind g y /\ SF y -> g y x0 /\ SF y
H7:X x0
H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2
x1:R
H8:f' x1 x0 /\ SF x1

ind g x1 /\ SF x1 -> False
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> Prop
H2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, ind g y /\ SF y -> g y x0 /\ SF y
H7:X x0
H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2
x1:R
H8:f' x1 x0 /\ SF x1
ind g x1 /\ SF x1
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> Prop
H2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF y
X x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)
~ (exists y : R, intersection_family (subfamily g SF) y)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> Prop
H2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, ind g y /\ SF y -> g y x0 /\ SF y
H7:X x0
H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2
x1:R
H8:f' x1 x0 /\ SF x1

ind g x1 /\ SF x1
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> Prop
H2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF y
X x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)
~ (exists y : R, intersection_family (subfamily g SF) y)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> Prop
H2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, ind g y /\ SF y -> g y x0 /\ SF y
H7:X x0
H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2
x1:R
H8:f' x1 x0 /\ SF x1

ind g x1
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> Prop
H2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, ind g y /\ SF y -> g y x0 /\ SF y
H7:X x0
H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2
x1:R
H8:f' x1 x0 /\ SF x1
SF x1
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> Prop
H2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF y
X x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)
~ (exists y : R, intersection_family (subfamily g SF) y)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> Prop
H2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, ind g y /\ SF y -> g y x0 /\ SF y
H7:X x0
H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2
x1:R
H8:f' x1 x0 /\ SF x1

exists y : R, f0 x1 y
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> Prop
H2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, ind g y /\ SF y -> g y x0 /\ SF y
H7:X x0
H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2
x1:R
H8:f' x1 x0 /\ SF x1
SF x1
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> Prop
H2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF y
X x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)
~ (exists y : R, intersection_family (subfamily g SF) y)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> Prop
H2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, ind g y /\ SF y -> g y x0 /\ SF y
H7:X x0
H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2
x1:R
H8:f' x1 x0 /\ SF x1

SF x1
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> Prop
H2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF y
X x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)
~ (exists y : R, intersection_family (subfamily g SF) y)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x1 y : R => complementary (g x1) y /\ D' x1:R -> R -> Prop
H2:forall x1 : R, (exists y : R, f' x1 y) -> D' x1
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF y

X x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)
~ (exists y : R, intersection_family (subfamily g SF) y)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:forall x2 : R, (ind g x2 -> included (g x2) X) /\ ~ (exists y : R, intersection_family g y)
D':=ind g:R -> Prop
f':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> Prop
H2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF y
x1:R
H7:intersection_domain (ind g) SF x1
H8:g x1 x0 /\ SF x1
H9:g x1 x0
H10:SF x1

ind g x1 -> X x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:forall x2 : R, (ind g x2 -> included (g x2) X) /\ ~ (exists y : R, intersection_family g y)
D':=ind g:R -> Prop
f':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> Prop
H2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF y
x1:R
H7:intersection_domain (ind g) SF x1
H8:g x1 x0 /\ SF x1
H9:g x1 x0
H10:SF x1
ind g x1
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)
~ (exists y : R, intersection_family (subfamily g SF) y)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:forall x2 : R, (ind g x2 -> included (g x2) X) /\ ~ (exists y : R, intersection_family g y)
D':=ind g:R -> Prop
f':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> Prop
H2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF y
x1:R
H7:intersection_domain (ind g) SF x1
H8:g x1 x0 /\ SF x1
H9:g x1 x0
H10:SF x1
H11:ind g x1
H12:ind g x1 -> included (g x1) X
H13:~ (exists y : R, intersection_family g y)

ind g x1
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:forall x2 : R, (ind g x2 -> included (g x2) X) /\ ~ (exists y : R, intersection_family g y)
D':=ind g:R -> Prop
f':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> Prop
H2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF y
x1:R
H7:intersection_domain (ind g) SF x1
H8:g x1 x0 /\ SF x1
H9:g x1 x0
H10:SF x1
H11:ind g x1
H12:ind g x1 -> included (g x1) X
H13:~ (exists y : R, intersection_family g y)
g x1 x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:forall x2 : R, (ind g x2 -> included (g x2) X) /\ ~ (exists y : R, intersection_family g y)
D':=ind g:R -> Prop
f':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> Prop
H2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF y
x1:R
H7:intersection_domain (ind g) SF x1
H8:g x1 x0 /\ SF x1
H9:g x1 x0
H10:SF x1
ind g x1
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)
~ (exists y : R, intersection_family (subfamily g SF) y)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:forall x2 : R, (ind g x2 -> included (g x2) X) /\ ~ (exists y : R, intersection_family g y)
D':=ind g:R -> Prop
f':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> Prop
H2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF y
x1:R
H7:intersection_domain (ind g) SF x1
H8:g x1 x0 /\ SF x1
H9:g x1 x0
H10:SF x1
H11:ind g x1
H12:ind g x1 -> included (g x1) X
H13:~ (exists y : R, intersection_family g y)

g x1 x0
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:forall x2 : R, (ind g x2 -> included (g x2) X) /\ ~ (exists y : R, intersection_family g y)
D':=ind g:R -> Prop
f':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> Prop
H2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF y
x1:R
H7:intersection_domain (ind g) SF x1
H8:g x1 x0 /\ SF x1
H9:g x1 x0
H10:SF x1
ind g x1
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)
~ (exists y : R, intersection_family (subfamily g SF) y)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:forall x2 : R, (ind g x2 -> included (g x2) X) /\ ~ (exists y : R, intersection_family g y)
D':=ind g:R -> Prop
f':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> Prop
H2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':exists y : R, intersection_domain (ind g) SF y
H5:exists y : R, intersection_family (subfamily g SF) y
x0:R
H6:forall y : R, intersection_domain (ind g) SF y -> g y x0 /\ SF y
x1:R
H7:intersection_domain (ind g) SF x1
H8:g x1 x0 /\ SF x1
H9:g x1 x0
H10:SF x1

ind g x1
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)
~ (exists y : R, intersection_family (subfamily g SF) y)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)

~ (exists y : R, intersection_family (subfamily g SF) y)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)
H4:covering X (subfamily f0 SF)

(exists z : R, X z) -> ~ (exists y : R, intersection_family (subfamily g SF) y)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)
H4:covering X (subfamily f0 SF)
exists z : R, X z
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> Prop
H2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)
H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2
x0:R
H5:X x0
x1:R
H6:f' x1 x0 /\ SF x1
H7:f' x1 x0
H8:SF x1

ind g x1
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> Prop
H2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)
H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2
x0:R
H5:X x0
x1:R
H6:f' x1 x0 /\ SF x1
H7:f' x1 x0
H8:SF x1
SF x1
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)
H4:covering X (subfamily f0 SF)
exists z : R, X z
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x2 y : R => complementary (g x2) y /\ D' x2:R -> R -> Prop
H2:forall x2 : R, (exists y : R, f' x2 y) -> D' x2
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)
H4:forall x2 : R, X x2 -> exists y : R, subfamily f0 SF y x2
x0:R
H5:X x0
x1:R
H6:f' x1 x0 /\ SF x1
H7:f' x1 x0
H8:SF x1

SF x1
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)
H4:covering X (subfamily f0 SF)
exists z : R, X z
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x0 y : R => complementary (g x0) y /\ D' x0:R -> R -> Prop
H2:forall x0 : R, (exists y : R, f' x0 y) -> D' x0
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
x:R
Hyp':~ (exists y : R, intersection_domain (ind g) SF y)
H4:covering X (subfamily f0 SF)

exists z : R, X z
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)
family_finite (subfamily g SF)
X:R -> Prop
H:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
Hyp:exists z : R, X z
g:family
H0:family_closed_set g
H1:intersection_vide_in X g
D':=ind g:R -> Prop
f':=fun x y : R => complementary (g x) y /\ D' x:R -> R -> Prop
H2:forall x : R, (exists y : R, f' x y) -> D' x
f0:={| ind := D'; f := f'; cond_fam := H2 |}:family
H3:covering_open_set X f0
SF:R -> Prop
H4:covering_finite X (subfamily f0 SF)

family_finite (subfamily g SF)
unfold covering_finite in H4; elim H4; clear H4; intros; unfold family_finite in H5; unfold domain_finite in H5; unfold family_finite; unfold domain_finite; elim H5; clear H5; intros l H5; exists l; intro; elim (H5 x); intros; split; intro; [ apply H6; simpl; simpl in H8; apply H8 | apply (H7 H8) ]. Qed.

forall (un : nat -> R) (X : R -> Prop), compact X -> (forall n : nat, X (un n)) -> exists l : R, ValAdh un l

forall (un : nat -> R) (X : R -> Prop), compact X -> (forall n : nat, X (un n)) -> exists l : R, ValAdh un l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)

(exists l : R, ValAdh_un un l) -> exists l : R, ValAdh un l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
exists l : R, ValAdh_un un l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)

exists l : R, ValAdh_un un l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)

exists z : R, X z
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
exists l : R, ValAdh_un un l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z

exists l : R, ValAdh_un un l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop

exists l : R, ValAdh_un un l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop

exists l : R, ValAdh_un un l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop

forall x : R, (exists y : R, g x y) -> D x
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
exists l : R, ValAdh_un un l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x1 : R => exists n : nat, x1 = INR n:R -> Prop
g:=fun x1 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x1 <= INR p) /\ D x1):R -> R -> Prop
x:R
H2:exists y : R, g x y
x0:R
H3:forall V : R -> Prop, neighbourhood V x0 -> exists y : R, intersection_domain V (fun y0 : R => (exists p : nat, y0 = un p /\ x <= INR p) /\ D x) y

D x
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
exists l : R, ValAdh_un un l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x1 : R => exists n : nat, x1 = INR n:R -> Prop
g:=fun x1 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x1 <= INR p) /\ D x1):R -> R -> Prop
x:R
H2:exists y : R, g x y
x0:R
H3:forall V : R -> Prop, neighbourhood V x0 -> exists y : R, intersection_domain V (fun y0 : R => (exists p : nat, y0 = un p /\ x <= INR p) /\ D x) y

neighbourhood (disc x0 {| pos := 1; cond_pos := Rlt_0_1 |}) x0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x1 : R => exists n : nat, x1 = INR n:R -> Prop
g:=fun x1 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x1 <= INR p) /\ D x1):R -> R -> Prop
x:R
H2:exists y : R, g x y
x0:R
H3:forall V : R -> Prop, neighbourhood V x0 -> exists y : R, intersection_domain V (fun y0 : R => (exists p : nat, y0 = un p /\ x <= INR p) /\ D x) y
H4:neighbourhood (disc x0 {| pos := 1; cond_pos := Rlt_0_1 |}) x0
D x
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
exists l : R, ValAdh_un un l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x1 : R => exists n : nat, x1 = INR n:R -> Prop
g:=fun x1 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x1 <= INR p) /\ D x1):R -> R -> Prop
x:R
H2:exists y : R, g x y
x0:R
H3:forall V : R -> Prop, neighbourhood V x0 -> exists y : R, intersection_domain V (fun y0 : R => (exists p : nat, y0 = un p /\ x <= INR p) /\ D x) y
H4:neighbourhood (disc x0 {| pos := 1; cond_pos := Rlt_0_1 |}) x0

D x
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
exists l : R, ValAdh_un un l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x

exists l : R, ValAdh_un un l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family

exists l : R, ValAdh_un un l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)

exists l : R, ValAdh_un un l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:exists l : R, ValAdh_un un l

exists l : R, ValAdh_un un l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
exists l : R, ValAdh_un un l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)

exists l : R, ValAdh_un un l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)

family_closed_set f0 -> exists l : R, ValAdh_un un l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0

intersection_vide_in X f0 -> exists l : R, ValAdh_un un l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
intersection_vide_in X f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)

exists l : R, ValAdh_un un l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
intersection_vide_in X f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x : R, (ind (subfamily f0 SF) x -> included (subfamily f0 SF x) X) /\ ~ (exists y : R, intersection_family (subfamily f0 SF) y)
H10:~ (exists y : R, intersection_family (subfamily f0 SF) y)
l:Rlist
H9:forall x : R, ind (subfamily f0 SF) x <-> In x l
r:=MaxRlist l:R

D r -> exists y : R, intersection_family (subfamily f0 SF) y
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x : R, (ind (subfamily f0 SF) x -> included (subfamily f0 SF x) X) /\ ~ (exists y : R, intersection_family (subfamily f0 SF) y)
H10:~ (exists y : R, intersection_family (subfamily f0 SF) y)
l:Rlist
H9:forall x : R, ind (subfamily f0 SF) x <-> In x l
r:=MaxRlist l:R
D r
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
intersection_vide_in X f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x0 : R => exists n : nat, x0 = INR n:R -> Prop
g:=fun x0 : R => adherence (fun y0 : R => (exists p : nat, y0 = un p /\ x0 <= INR p) /\ D x0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> D x0
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x0 : R, (ind (subfamily f0 SF) x0 -> included (subfamily f0 SF x0) X) /\ ~ (exists y0 : R, intersection_family (subfamily f0 SF) y0)
H10:~ (exists y0 : R, intersection_family (subfamily f0 SF) y0)
l:Rlist
H9:forall x0 : R, ind (subfamily f0 SF) x0 <-> In x0 l
r:=MaxRlist l:R
H11:exists n : nat, r = INR n
x:nat
H12:r = INR x
y:R
H13:D y /\ SF y

g y (un x)
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x0 : R => exists n : nat, x0 = INR n:R -> Prop
g:=fun x0 : R => adherence (fun y0 : R => (exists p : nat, y0 = un p /\ x0 <= INR p) /\ D x0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> D x0
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x0 : R, (ind (subfamily f0 SF) x0 -> included (subfamily f0 SF x0) X) /\ ~ (exists y0 : R, intersection_family (subfamily f0 SF) y0)
H10:~ (exists y0 : R, intersection_family (subfamily f0 SF) y0)
l:Rlist
H9:forall x0 : R, ind (subfamily f0 SF) x0 <-> In x0 l
r:=MaxRlist l:R
H11:exists n : nat, r = INR n
x:nat
H12:r = INR x
y:R
H13:D y /\ SF y
SF y
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x : R, (ind (subfamily f0 SF) x -> included (subfamily f0 SF x) X) /\ ~ (exists y : R, intersection_family (subfamily f0 SF) y)
H10:~ (exists y : R, intersection_family (subfamily f0 SF) y)
l:Rlist
H9:forall x : R, ind (subfamily f0 SF) x <-> In x l
r:=MaxRlist l:R
D r
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
intersection_vide_in X f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x0 : R => exists n : nat, x0 = INR n:R -> Prop
g:=fun x0 : R => adherence (fun y0 : R => (exists p : nat, y0 = un p /\ x0 <= INR p) /\ D x0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> D x0
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x0 : R, (ind (subfamily f0 SF) x0 -> included (subfamily f0 SF x0) X) /\ ~ (exists y0 : R, intersection_family (subfamily f0 SF) y0)
H10:~ (exists y0 : R, intersection_family (subfamily f0 SF) y0)
l:Rlist
H9:forall x0 : R, ind (subfamily f0 SF) x0 <-> In x0 l
r:=MaxRlist l:R
H11:exists n : nat, r = INR n
x:nat
H12:r = INR x
y:R
H13:D y /\ SF y

exists p : nat, un x = un p /\ y <= INR p
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x0 : R => exists n : nat, x0 = INR n:R -> Prop
g:=fun x0 : R => adherence (fun y0 : R => (exists p : nat, y0 = un p /\ x0 <= INR p) /\ D x0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> D x0
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x0 : R, (ind (subfamily f0 SF) x0 -> included (subfamily f0 SF x0) X) /\ ~ (exists y0 : R, intersection_family (subfamily f0 SF) y0)
H10:~ (exists y0 : R, intersection_family (subfamily f0 SF) y0)
l:Rlist
H9:forall x0 : R, ind (subfamily f0 SF) x0 <-> In x0 l
r:=MaxRlist l:R
H11:exists n : nat, r = INR n
x:nat
H12:r = INR x
y:R
H13:D y /\ SF y
D y
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x0 : R => exists n : nat, x0 = INR n:R -> Prop
g:=fun x0 : R => adherence (fun y0 : R => (exists p : nat, y0 = un p /\ x0 <= INR p) /\ D x0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> D x0
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x0 : R, (ind (subfamily f0 SF) x0 -> included (subfamily f0 SF x0) X) /\ ~ (exists y0 : R, intersection_family (subfamily f0 SF) y0)
H10:~ (exists y0 : R, intersection_family (subfamily f0 SF) y0)
l:Rlist
H9:forall x0 : R, ind (subfamily f0 SF) x0 <-> In x0 l
r:=MaxRlist l:R
H11:exists n : nat, r = INR n
x:nat
H12:r = INR x
y:R
H13:D y /\ SF y
SF y
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x : R, (ind (subfamily f0 SF) x -> included (subfamily f0 SF x) X) /\ ~ (exists y : R, intersection_family (subfamily f0 SF) y)
H10:~ (exists y : R, intersection_family (subfamily f0 SF) y)
l:Rlist
H9:forall x : R, ind (subfamily f0 SF) x <-> In x l
r:=MaxRlist l:R
D r
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
intersection_vide_in X f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x0 : R => exists n : nat, x0 = INR n:R -> Prop
g:=fun x0 : R => adherence (fun y0 : R => (exists p : nat, y0 = un p /\ x0 <= INR p) /\ D x0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> D x0
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x0 : R, (ind (subfamily f0 SF) x0 -> included (subfamily f0 SF x0) X) /\ ~ (exists y0 : R, intersection_family (subfamily f0 SF) y0)
H10:~ (exists y0 : R, intersection_family (subfamily f0 SF) y0)
l:Rlist
H9:forall x0 : R, ind (subfamily f0 SF) x0 <-> In x0 l
r:=MaxRlist l:R
H11:exists n : nat, r = INR n
x:nat
H12:r = INR x
y:R
H13:D y /\ SF y

D y
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x0 : R => exists n : nat, x0 = INR n:R -> Prop
g:=fun x0 : R => adherence (fun y0 : R => (exists p : nat, y0 = un p /\ x0 <= INR p) /\ D x0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> D x0
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x0 : R, (ind (subfamily f0 SF) x0 -> included (subfamily f0 SF x0) X) /\ ~ (exists y0 : R, intersection_family (subfamily f0 SF) y0)
H10:~ (exists y0 : R, intersection_family (subfamily f0 SF) y0)
l:Rlist
H9:forall x0 : R, ind (subfamily f0 SF) x0 <-> In x0 l
r:=MaxRlist l:R
H11:exists n : nat, r = INR n
x:nat
H12:r = INR x
y:R
H13:D y /\ SF y
SF y
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x : R, (ind (subfamily f0 SF) x -> included (subfamily f0 SF x) X) /\ ~ (exists y : R, intersection_family (subfamily f0 SF) y)
H10:~ (exists y : R, intersection_family (subfamily f0 SF) y)
l:Rlist
H9:forall x : R, ind (subfamily f0 SF) x <-> In x l
r:=MaxRlist l:R
D r
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
intersection_vide_in X f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x0 : R => exists n : nat, x0 = INR n:R -> Prop
g:=fun x0 : R => adherence (fun y0 : R => (exists p : nat, y0 = un p /\ x0 <= INR p) /\ D x0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> D x0
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x0 : R, (ind (subfamily f0 SF) x0 -> included (subfamily f0 SF x0) X) /\ ~ (exists y0 : R, intersection_family (subfamily f0 SF) y0)
H10:~ (exists y0 : R, intersection_family (subfamily f0 SF) y0)
l:Rlist
H9:forall x0 : R, ind (subfamily f0 SF) x0 <-> In x0 l
r:=MaxRlist l:R
H11:exists n : nat, r = INR n
x:nat
H12:r = INR x
y:R
H13:D y /\ SF y

SF y
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x : R, (ind (subfamily f0 SF) x -> included (subfamily f0 SF x) X) /\ ~ (exists y : R, intersection_family (subfamily f0 SF) y)
H10:~ (exists y : R, intersection_family (subfamily f0 SF) y)
l:Rlist
H9:forall x : R, ind (subfamily f0 SF) x <-> In x l
r:=MaxRlist l:R
D r
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
intersection_vide_in X f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x : R, (ind (subfamily f0 SF) x -> included (subfamily f0 SF x) X) /\ ~ (exists y : R, intersection_family (subfamily f0 SF) y)
H10:~ (exists y : R, intersection_family (subfamily f0 SF) y)
l:Rlist
H9:forall x : R, ind (subfamily f0 SF) x <-> In x l
r:=MaxRlist l:R

D r
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
intersection_vide_in X f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x : R, (ind (subfamily f0 SF) x -> included (subfamily f0 SF x) X) /\ ~ (exists y : R, intersection_family (subfamily f0 SF) y)
H10:~ (exists y : R, intersection_family (subfamily f0 SF) y)
l:Rlist
H9:forall x : R, ind (subfamily f0 SF) x <-> In x l
r:=MaxRlist l:R
H11:ind (subfamily f0 SF) r -> In r l
H12:In r l -> ind (subfamily f0 SF) r

D r
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
intersection_vide_in X f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x : R, (ind (subfamily f0 SF) x -> included (subfamily f0 SF x) X) /\ ~ (exists y : R, intersection_family (subfamily f0 SF) y)
H10:~ (exists y : R, intersection_family (subfamily f0 SF) y)
l:Rlist
H9:forall x : R, ind (subfamily f0 SF) x <-> In x l
r:=MaxRlist l:R
H11:ind (subfamily f0 SF) r -> In r l
H12:In r l -> D r /\ SF r

In r l -> D r
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x : R, (ind (subfamily f0 SF) x -> included (subfamily f0 SF x) X) /\ ~ (exists y : R, intersection_family (subfamily f0 SF) y)
H10:~ (exists y : R, intersection_family (subfamily f0 SF) y)
l:Rlist
H9:forall x : R, ind (subfamily f0 SF) x <-> In x l
r:=MaxRlist l:R
H11:ind (subfamily f0 SF) r -> In r l
H12:In r l -> D r /\ SF r
In r l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
intersection_vide_in X f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x : R, (ind (subfamily f0 SF) x -> included (subfamily f0 SF x) X) /\ ~ (exists y : R, intersection_family (subfamily f0 SF) y)
H10:~ (exists y : R, intersection_family (subfamily f0 SF) y)
l:Rlist
H9:forall x : R, ind (subfamily f0 SF) x <-> In x l
r:=MaxRlist l:R
H11:ind (subfamily f0 SF) r -> In r l
H12:In r l -> D r /\ SF r

In r l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
intersection_vide_in X f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x : R, (ind (subfamily f0 SF) x -> included (subfamily f0 SF x) X) /\ ~ (exists y : R, intersection_family (subfamily f0 SF) y)
H10:~ (exists y : R, intersection_family (subfamily f0 SF) y)
l:Rlist
H9:forall x : R, ind (subfamily f0 SF) x <-> In x l
r:=MaxRlist l:R
H11:ind (subfamily f0 SF) r -> In r l
H12:In r l -> D r /\ SF r

(exists z : R, intersection_domain (ind f0) SF z) -> exists y : R, In y l
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x : R, (ind (subfamily f0 SF) x -> included (subfamily f0 SF x) X) /\ ~ (exists y : R, intersection_family (subfamily f0 SF) y)
H10:~ (exists y : R, intersection_family (subfamily f0 SF) y)
l:Rlist
H9:forall x : R, ind (subfamily f0 SF) x <-> In x l
r:=MaxRlist l:R
H11:ind (subfamily f0 SF) r -> In r l
H12:In r l -> D r /\ SF r
exists z : R, intersection_domain (ind f0) SF z
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
intersection_vide_in X f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x : R, (ind (subfamily f0 SF) x -> included (subfamily f0 SF x) X) /\ ~ (exists y : R, intersection_family (subfamily f0 SF) y)
H10:~ (exists y : R, intersection_family (subfamily f0 SF) y)
l:Rlist
H9:forall x : R, ind (subfamily f0 SF) x <-> In x l
r:=MaxRlist l:R
H11:ind (subfamily f0 SF) r -> In r l
H12:In r l -> D r /\ SF r

exists z : R, intersection_domain (ind f0) SF z
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
intersection_vide_in X f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x : R, (ind (subfamily f0 SF) x -> included (subfamily f0 SF x) X) /\ ~ (exists y : R, intersection_family (subfamily f0 SF) y)
H10:~ (exists y : R, intersection_family (subfamily f0 SF) y)
l:Rlist
H9:forall x : R, ind (subfamily f0 SF) x <-> In x l
r:=MaxRlist l:R
H11:ind (subfamily f0 SF) r -> In r l
H12:In r l -> D r /\ SF r
H13:exists z : R, intersection_domain (ind f0) SF z

exists z : R, intersection_domain (ind f0) SF z
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x : R, (ind (subfamily f0 SF) x -> included (subfamily f0 SF x) X) /\ ~ (exists y : R, intersection_family (subfamily f0 SF) y)
H10:~ (exists y : R, intersection_family (subfamily f0 SF) y)
l:Rlist
H9:forall x : R, ind (subfamily f0 SF) x <-> In x l
r:=MaxRlist l:R
H11:ind (subfamily f0 SF) r -> In r l
H12:In r l -> D r /\ SF r
H13:~ (exists z : R, intersection_domain (ind f0) SF z)
exists z : R, intersection_domain (ind f0) SF z
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
intersection_vide_in X f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l0 : R, ValAdh_un un l0)
H5:family_closed_set f0
H6:intersection_vide_in X f0
H7:exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
SF:R -> Prop
H8:forall x : R, (ind (subfamily f0 SF) x -> included (subfamily f0 SF x) X) /\ ~ (exists y : R, intersection_family (subfamily f0 SF) y)
H10:~ (exists y : R, intersection_family (subfamily f0 SF) y)
l:Rlist
H9:forall x : R, ind (subfamily f0 SF) x <-> In x l
r:=MaxRlist l:R
H11:ind (subfamily f0 SF) r -> In r l
H12:In r l -> D r /\ SF r
H13:~ (exists z : R, intersection_domain (ind f0) SF z)

exists z : R, intersection_domain (ind f0) SF z
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
intersection_vide_in X f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0

intersection_vide_in X f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x0 : R => exists n : nat, x0 = INR n:R -> Prop
g:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
x:R

ind f0 x -> included (f0 x) X
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x0 : R => exists n : nat, x0 = INR n:R -> Prop
g:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
x:R
~ (exists y : R, intersection_family f0 y)
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x0 : R => exists n : nat, x0 = INR n:R -> Prop
g:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
x:R
H6:D x

included (adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x)) (adherence X)
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x0 : R => exists n : nat, x0 = INR n:R -> Prop
g:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
x:R
H6:D x
included (adherence X) X
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x0 : R => exists n : nat, x0 = INR n:R -> Prop
g:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
x:R
~ (exists y : R, intersection_family f0 y)
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x0 : R => exists n : nat, x0 = INR n:R -> Prop
g:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
x:R
H6:D x

included (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x) X
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x0 : R => exists n : nat, x0 = INR n:R -> Prop
g:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
x:R
H6:D x
included (adherence X) X
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x0 : R => exists n : nat, x0 = INR n:R -> Prop
g:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
x:R
~ (exists y : R, intersection_family f0 y)
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x0 : R => exists n : nat, x0 = INR n:R -> Prop
g:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
x:R
H6:D x

included (adherence X) X
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x0 : R => exists n : nat, x0 = INR n:R -> Prop
g:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
x:R
~ (exists y : R, intersection_family f0 y)
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x0 : R => exists n : nat, x0 = INR n:R -> Prop
g:=fun x0 : R => adherence (fun y : R => (exists p : nat, y = un p /\ x0 <= INR p) /\ D x0):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> D x0
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
H5:family_closed_set f0
x:R

~ (exists y : R, intersection_family f0 y)
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)
family_closed_set f0
un:nat -> R
X:R -> Prop
H:compact X
H0:forall n : nat, X (un n)
H1:exists z : R, X z
D:=fun x : R => exists n : nat, x = INR n:R -> Prop
g:=fun x : R => adherence (fun y : R => (exists p : nat, y = un p /\ x <= INR p) /\ D x):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> D x
f0:={| ind := D; f := g; cond_fam := H2 |}:family
H3:family_closed_set f0 -> intersection_vide_in X f0 -> exists D0 : R -> Prop, intersection_vide_finite_in X (subfamily f0 D0)
H4:~ (exists l : R, ValAdh_un un l)

family_closed_set f0
unfold family_closed_set; unfold f0; simpl; unfold g; intro; apply adherence_P3. Qed. (********************************************************)

Proof of Heine's theorem

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

Definition uniform_continuity (f:R -> R) (X:R -> Prop) : Prop :=
  forall eps:posreal,
    exists delta : posreal,
      (forall x y:R,
        X x -> X y -> Rabs (x - y) < delta -> Rabs (f x - f y) < eps).


forall (E : R -> Prop) (x y : R), is_lub E x -> is_lub E y -> x = y

forall (E : R -> Prop) (x y : R), is_lub E x -> is_lub E y -> x = y
unfold is_lub; intros; elim H; elim H0; intros; apply Rle_antisym; [ apply (H4 _ H1) | apply (H2 _ H3) ]. Qed.

forall X : R -> Prop, ~ (exists y : R, X y) \/ (exists y : R, X y /\ (forall x : R, X x -> x = y)) \/ (exists x y : R, X x /\ X y /\ x <> y)

forall X : R -> Prop, ~ (exists y : R, X y) \/ (exists y : R, X y /\ (forall x : R, X x -> x = y)) \/ (exists x y : R, X x /\ X y /\ x <> y)
X:R -> Prop
H:exists y : R, X y

~ (exists y : R, X y) \/ (exists y : R, X y /\ (forall x : R, X x -> x = y)) \/ (exists x y : R, X x /\ X y /\ x <> y)
X:R -> Prop
H:~ (exists y : R, X y)
~ (exists y : R, X y) \/ (exists y : R, X y /\ (forall x : R, X x -> x = y)) \/ (exists x y : R, X x /\ X y /\ x <> y)
X:R -> Prop
H:exists y : R, X y
x:R
H0:X x
H1:exists y : R, X y /\ y <> x

(exists y : R, X y /\ (forall x0 : R, X x0 -> x0 = y)) \/ (exists x0 y : R, X x0 /\ X y /\ x0 <> y)
X:R -> Prop
H:exists y : R, X y
x:R
H0:X x
H1:~ (exists y : R, X y /\ y <> x)
(exists y : R, X y /\ (forall x0 : R, X x0 -> x0 = y)) \/ (exists x0 y : R, X x0 /\ X y /\ x0 <> y)
X:R -> Prop
H:~ (exists y : R, X y)
~ (exists y : R, X y) \/ (exists y : R, X y /\ (forall x : R, X x -> x = y)) \/ (exists x y : R, X x /\ X y /\ x <> y)
X:R -> Prop
H:exists y : R, X y
x:R
H0:X x
H1:exists y : R, X y /\ y <> x
x0:R
H2:X x0 /\ x0 <> x
H3:X x0
H4:x0 <> x

X x /\ X x0 /\ x <> x0
X:R -> Prop
H:exists y : R, X y
x:R
H0:X x
H1:~ (exists y : R, X y /\ y <> x)
(exists y : R, X y /\ (forall x0 : R, X x0 -> x0 = y)) \/ (exists x0 y : R, X x0 /\ X y /\ x0 <> y)
X:R -> Prop
H:~ (exists y : R, X y)
~ (exists y : R, X y) \/ (exists y : R, X y /\ (forall x : R, X x -> x = y)) \/ (exists x y : R, X x /\ X y /\ x <> y)
X:R -> Prop
H:exists y : R, X y
x:R
H0:X x
H1:~ (exists y : R, X y /\ y <> x)

(exists y : R, X y /\ (forall x0 : R, X x0 -> x0 = y)) \/ (exists x0 y : R, X x0 /\ X y /\ x0 <> y)
X:R -> Prop
H:~ (exists y : R, X y)
~ (exists y : R, X y) \/ (exists y : R, X y /\ (forall x : R, X x -> x = y)) \/ (exists x y : R, X x /\ X y /\ x <> y)
X:R -> Prop
H:exists y : R, X y
x:R
H0:X x
H1:~ (exists y : R, X y /\ y <> x)

X x
X:R -> Prop
H:exists y : R, X y
x:R
H0:X x
H1:~ (exists y : R, X y /\ y <> x)
forall x0 : R, X x0 -> x0 = x
X:R -> Prop
H:~ (exists y : R, X y)
~ (exists y : R, X y) \/ (exists y : R, X y /\ (forall x : R, X x -> x = y)) \/ (exists x y : R, X x /\ X y /\ x <> y)
X:R -> Prop
H:exists y : R, X y
x:R
H0:X x
H1:~ (exists y : R, X y /\ y <> x)

forall x0 : R, X x0 -> x0 = x
X:R -> Prop
H:~ (exists y : R, X y)
~ (exists y : R, X y) \/ (exists y : R, X y /\ (forall x : R, X x -> x = y)) \/ (exists x y : R, X x /\ X y /\ x <> y)
X:R -> Prop
H:exists y : R, X y
x:R
H0:X x
H1:~ (exists y : R, X y /\ y <> x)
x0:R
H2:X x0
H3:x0 = x

x0 = x
X:R -> Prop
H:exists y : R, X y
x:R
H0:X x
H1:~ (exists y : R, X y /\ y <> x)
x0:R
H2:X x0
H3:x0 <> x
x0 = x
X:R -> Prop
H:~ (exists y : R, X y)
~ (exists y : R, X y) \/ (exists y : R, X y /\ (forall x : R, X x -> x = y)) \/ (exists x y : R, X x /\ X y /\ x <> y)
X:R -> Prop
H:exists y : R, X y
x:R
H0:X x
H1:~ (exists y : R, X y /\ y <> x)
x0:R
H2:X x0
H3:x0 <> x

x0 = x
X:R -> Prop
H:~ (exists y : R, X y)
~ (exists y : R, X y) \/ (exists y : R, X y /\ (forall x : R, X x -> x = y)) \/ (exists x y : R, X x /\ X y /\ x <> y)
X:R -> Prop
H:~ (exists y : R, X y)

~ (exists y : R, X y) \/ (exists y : R, X y /\ (forall x : R, X x -> x = y)) \/ (exists x y : R, X x /\ X y /\ x <> y)
left; assumption. Qed.

forall (f0 : R -> R) (X : R -> Prop), compact X -> (forall x : R, X x -> continuity_pt f0 x) -> uniform_continuity f0 X

forall (f0 : R -> R) (X : R -> Prop), compact X -> (forall x : R, X x -> continuity_pt f0 x) -> uniform_continuity f0 X
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x : R, X x -> continuity_pt f0 x
Hyp:~ (exists y : R, X y)

uniform_continuity f0 X
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x : R, X x -> continuity_pt f0 x
Hyp:(exists y : R, X y /\ (forall x : R, X x -> x = y)) \/ (exists x y : R, X x /\ X y /\ x <> y)
uniform_continuity f0 X
(* X is empty *)
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x : R, X x -> continuity_pt f0 x
Hyp:(exists y : R, X y /\ (forall x : R, X x -> x = y)) \/ (exists x y : R, X x /\ X y /\ x <> y)

uniform_continuity f0 X
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x : R, X x -> continuity_pt f0 x
Hyp:exists y : R, X y /\ (forall x : R, X x -> x = y)

uniform_continuity f0 X
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x : R, X x -> continuity_pt f0 x
Hyp:exists x y : R, X x /\ X y /\ x <> y
uniform_continuity f0 X
(* X has only one element *)
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x : R, X x -> continuity_pt f0 x
Hyp:exists x y : R, X x /\ X y /\ x <> y

uniform_continuity f0 X
(* X has at least two distinct elements *)
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x : R, X x -> continuity_pt f0 x
Hyp:exists x y : R, X x /\ X y /\ x <> y

exists m M : R, (forall x : R, X x -> m <= x <= M) /\ m < M
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x : R, X x -> continuity_pt f0 x
Hyp:exists x y : R, X x /\ X y /\ x <> y
X_enc:exists m M : R, (forall x : R, X x -> m <= x <= M) /\ m < M
uniform_continuity f0 X
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x1 : R, X x1 -> continuity_pt f0 x1
Hyp:exists x1 y : R, X x1 /\ X y /\ x1 <> y
H1:exists m M : R, forall x1 : R, X x1 -> m <= x1 <= M
x:R
H2:exists M : R, forall x1 : R, X x1 -> x <= x1 <= M
x0:R
H3:forall x1 : R, X x1 -> x <= x1 <= x0

forall x1 : R, X x1 -> x <= x1 <= x0
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x1 : R, X x1 -> continuity_pt f0 x1
Hyp:exists x1 y : R, X x1 /\ X y /\ x1 <> y
H1:exists m M : R, forall x1 : R, X x1 -> m <= x1 <= M
x:R
H2:exists M : R, forall x1 : R, X x1 -> x <= x1 <= M
x0:R
H3:forall x1 : R, X x1 -> x <= x1 <= x0
x < x0
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x : R, X x -> continuity_pt f0 x
Hyp:exists x y : R, X x /\ X y /\ x <> y
X_enc:exists m M : R, (forall x : R, X x -> m <= x <= M) /\ m < M
uniform_continuity f0 X
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x1 : R, X x1 -> continuity_pt f0 x1
Hyp:exists x1 y : R, X x1 /\ X y /\ x1 <> y
H1:exists m M : R, forall x1 : R, X x1 -> m <= x1 <= M
x:R
H2:exists M : R, forall x1 : R, X x1 -> x <= x1 <= M
x0:R
H3:forall x1 : R, X x1 -> x <= x1 <= x0

x < x0
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x : R, X x -> continuity_pt f0 x
Hyp:exists x y : R, X x /\ X y /\ x <> y
X_enc:exists m M : R, (forall x : R, X x -> m <= x <= M) /\ m < M
uniform_continuity f0 X
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x3 : R, X x3 -> continuity_pt f0 x3
Hyp:exists x3 y : R, X x3 /\ X y /\ x3 <> y
H1:exists m M : R, forall x3 : R, X x3 -> m <= x3 <= M
x:R
H2:exists M : R, forall x3 : R, X x3 -> x <= x3 <= M
x0:R
H3:forall x3 : R, X x3 -> x <= x3 <= x0
x1:R
H4:exists y : R, X x1 /\ X y /\ x1 <> y
x2:R
H5:X x1 /\ X x2 /\ x1 <> x2
H6:X x1
H8:X x2
H9:x1 <> x2
H10:x <= x1 <= x0
H11:x <= x2 <= x0
H7:x <= x1
H12:x1 <= x0
H13:x <= x2
H14:x2 <= x0
r:x < x0

x < x0
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x3 : R, X x3 -> continuity_pt f0 x3
Hyp:exists x3 y : R, X x3 /\ X y /\ x3 <> y
H1:exists m M : R, forall x3 : R, X x3 -> m <= x3 <= M
x:R
H2:exists M : R, forall x3 : R, X x3 -> x <= x3 <= M
x0:R
H3:forall x3 : R, X x3 -> x <= x3 <= x0
x1:R
H4:exists y : R, X x1 /\ X y /\ x1 <> y
x2:R
H5:X x1 /\ X x2 /\ x1 <> x2
H6:X x1
H8:X x2
H9:x1 <> x2
H10:x <= x1 <= x0
H11:x <= x2 <= x0
H7:x <= x1
H12:x1 <= x0
H13:x <= x2
H14:x2 <= x0
H15:x = x0
x < x0
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x3 : R, X x3 -> continuity_pt f0 x3
Hyp:exists x3 y : R, X x3 /\ X y /\ x3 <> y
H1:exists m M : R, forall x3 : R, X x3 -> m <= x3 <= M
x:R
H2:exists M : R, forall x3 : R, X x3 -> x <= x3 <= M
x0:R
H3:forall x3 : R, X x3 -> x <= x3 <= x0
x1:R
H4:exists y : R, X x1 /\ X y /\ x1 <> y
x2:R
H5:X x1 /\ X x2 /\ x1 <> x2
H6:X x1
H8:X x2
H9:x1 <> x2
H10:x <= x1 <= x0
H11:x <= x2 <= x0
H7:x <= x1
H12:x1 <= x0
H13:x <= x2
H14:x2 <= x0
H15:x > x0
x < x0
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x : R, X x -> continuity_pt f0 x
Hyp:exists x y : R, X x /\ X y /\ x <> y
X_enc:exists m M : R, (forall x : R, X x -> m <= x <= M) /\ m < M
uniform_continuity f0 X
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x3 : R, X x3 -> continuity_pt f0 x3
Hyp:exists x3 y : R, X x3 /\ X y /\ x3 <> y
H1:exists m M : R, forall x3 : R, X x3 -> m <= x3 <= M
x:R
H2:exists M : R, forall x3 : R, X x3 -> x <= x3 <= M
x0:R
H3:forall x3 : R, X x3 -> x <= x3 <= x0
x1:R
H4:exists y : R, X x1 /\ X y /\ x1 <> y
x2:R
H5:X x1 /\ X x2 /\ x1 <> x2
H6:X x1
H8:X x2
H9:x1 <> x2
H10:x <= x1 <= x0
H11:x <= x2 <= x0
H7:x <= x1
H12:x1 <= x0
H13:x <= x2
H14:x2 <= x0
H15:x = x0

x < x0
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x3 : R, X x3 -> continuity_pt f0 x3
Hyp:exists x3 y : R, X x3 /\ X y /\ x3 <> y
H1:exists m M : R, forall x3 : R, X x3 -> m <= x3 <= M
x:R
H2:exists M : R, forall x3 : R, X x3 -> x <= x3 <= M
x0:R
H3:forall x3 : R, X x3 -> x <= x3 <= x0
x1:R
H4:exists y : R, X x1 /\ X y /\ x1 <> y
x2:R
H5:X x1 /\ X x2 /\ x1 <> x2
H6:X x1
H8:X x2
H9:x1 <> x2
H10:x <= x1 <= x0
H11:x <= x2 <= x0
H7:x <= x1
H12:x1 <= x0
H13:x <= x2
H14:x2 <= x0
H15:x > x0
x < x0
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x : R, X x -> continuity_pt f0 x
Hyp:exists x y : R, X x /\ X y /\ x <> y
X_enc:exists m M : R, (forall x : R, X x -> m <= x <= M) /\ m < M
uniform_continuity f0 X
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x3 : R, X x3 -> continuity_pt f0 x3
Hyp:exists x3 y : R, X x3 /\ X y /\ x3 <> y
H1:exists m M : R, forall x3 : R, X x3 -> m <= x3 <= M
x:R
H2:exists M : R, forall x3 : R, X x3 -> x <= x3 <= M
x0:R
H3:forall x3 : R, X x3 -> x <= x3 <= x0
x1:R
H4:exists y : R, X x1 /\ X y /\ x1 <> y
x2:R
H5:X x1 /\ X x2 /\ x1 <> x2
H6:X x1
H8:X x2
H9:x1 <> x2
H10:x <= x1 <= x0
H11:x <= x2 <= x0
H7:x <= x1
H12:x1 <= x0
H13:x <= x2
H14:x2 <= x0
H15:x > x0

x < x0
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x : R, X x -> continuity_pt f0 x
Hyp:exists x y : R, X x /\ X y /\ x <> y
X_enc:exists m M : R, (forall x : R, X x -> m <= x <= M) /\ m < M
uniform_continuity f0 X
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x : R, X x -> continuity_pt f0 x
Hyp:exists x y : R, X x /\ X y /\ x <> y
X_enc:exists m M : R, (forall x : R, X x -> m <= x <= M) /\ m < M

uniform_continuity f0 X
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal

forall t : posreal, 0 < t / 2
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2

exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop

exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop

forall x : R, (exists y : R, g x y) -> X x
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:compact X
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x

exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family

covering_open_set X f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family

covering X f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x

X x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x

exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop

bound E
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub E del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E

exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub E del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E

exists x1 : R, E x1
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x1 : R, E x1
exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub E del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H5:x0 > 0
H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2

0 < Rmin x0 (M - m) <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H5:x0 > 0
H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2
forall z : R, Rabs (z - x) < Rmin x0 (M - m) -> Rabs (f0 z - f0 x) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x1 : R, E x1
exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub E del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H5:x0 > 0
H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2

0 < Rmin x0 (M - m)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H5:x0 > 0
H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2
Rmin x0 (M - m) <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H5:x0 > 0
H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2
forall z : R, Rabs (z - x) < Rmin x0 (M - m) -> Rabs (f0 z - f0 x) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x1 : R, E x1
exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub E del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H5:x0 > 0
H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2
r:x0 <= M - m

0 < x0
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H5:x0 > 0
H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2
n:~ x0 <= M - m
0 < M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H5:x0 > 0
H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2
Rmin x0 (M - m) <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H5:x0 > 0
H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2
forall z : R, Rabs (z - x) < Rmin x0 (M - m) -> Rabs (f0 z - f0 x) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x1 : R, E x1
exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub E del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H5:x0 > 0
H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2
n:~ x0 <= M - m

0 < M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H5:x0 > 0
H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2
Rmin x0 (M - m) <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H5:x0 > 0
H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2
forall z : R, Rabs (z - x) < Rmin x0 (M - m) -> Rabs (f0 z - f0 x) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x1 : R, E x1
exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub E del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H5:x0 > 0
H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2

Rmin x0 (M - m) <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H5:x0 > 0
H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2
forall z : R, Rabs (z - x) < Rmin x0 (M - m) -> Rabs (f0 z - f0 x) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x1 : R, E x1
exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub E del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H5:x0 > 0
H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2

forall z : R, Rabs (z - x) < Rmin x0 (M - m) -> Rabs (f0 z - f0 x) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x1 : R, E x1
exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub E del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x1) < del -> Rabs (f0 z0 - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x1) < zeta -> Rabs (f0 z0 - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H6:bound E
H5:x0 > 0
H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2
z:R
H8:Rabs (z - x) < Rmin x0 (M - m)
H9:x = z

Rabs (f0 z - f0 x) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x1) < del -> Rabs (f0 z0 - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x1) < zeta -> Rabs (f0 z0 - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H6:bound E
H5:x0 > 0
H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2
z:R
H8:Rabs (z - x) < Rmin x0 (M - m)
H9:x <> z
Rabs (f0 z - f0 x) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x1 : R, E x1
exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub E del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x1) < del -> Rabs (f0 z0 - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x1) < zeta -> Rabs (f0 z0 - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H6:bound E
H5:x0 > 0
H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2
z:R
H8:Rabs (z - x) < Rmin x0 (M - m)
H9:x <> z

Rabs (f0 z - f0 x) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x1 : R, E x1
exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub E del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x1) < del -> Rabs (f0 z0 - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x1) < zeta -> Rabs (f0 z0 - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H6:bound E
H5:x0 > 0
H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2
z:R
H8:Rabs (z - x) < Rmin x0 (M - m)
H9:x <> z

D_x no_cond x z
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x1) < del -> Rabs (f0 z0 - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x1) < zeta -> Rabs (f0 z0 - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H6:bound E
H5:x0 > 0
H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2
z:R
H8:Rabs (z - x) < Rmin x0 (M - m)
H9:x <> z
Rabs (z - x) < x0
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x1 : R, E x1
exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub E del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x1) < del -> Rabs (f0 z0 - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x1) < zeta -> Rabs (f0 z0 - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H6:bound E
H5:x0 > 0
H7:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2
z:R
H8:Rabs (z - x) < Rmin x0 (M - m)
H9:x <> z

Rabs (z - x) < x0
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x1 : R, E x1
exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub E del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x1 : R, E x1

exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub E del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1

exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub E del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1

0 < x1 <= M - m -> exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub E del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
0 < x1 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m

forall z : R, Rabs (z - x) < {| pos := x1; cond_pos := H8 |} -> Rabs (f0 z - f0 x) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m
is_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
0 < x1 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x2) < del -> Rabs (f0 z0 - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x2) < zeta -> Rabs (f0 z0 - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m
z:R
H10:Rabs (z - x) < {| pos := x1; cond_pos := H8 |}

(exists alp : R, Rabs (z - x) < alp <= x1 /\ E alp) -> Rabs (f0 z - f0 x) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x2) < del -> Rabs (f0 z0 - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x2) < zeta -> Rabs (f0 z0 - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m
z:R
H10:Rabs (z - x) < {| pos := x1; cond_pos := H8 |}
exists alp : R, Rabs (z - x) < alp <= x1 /\ E alp
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m
is_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
0 < x1 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x3) < del -> Rabs (f0 z0 - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x3) < zeta -> Rabs (f0 z0 - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x3 : R, D_x no_cond x x3 /\ Rabs (x3 - x) < alp -> Rabs (f0 x3 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x3 : R, D_x no_cond x x3 /\ Rabs (x3 - x) < x0 -> Rabs (f0 x3 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x3 : R, E x3
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m
z:R
H10:Rabs (z - x) < {| pos := x1; cond_pos := H8 |}
H11:exists alp : R, Rabs (z - x) < alp <= x1 /\ E alp
x2:R
H12:Rabs (z - x) < x2 <= x1
H13:0 < x2 <= M - m /\ (forall z0 : R, Rabs (z0 - x) < x2 -> Rabs (f0 z0 - f0 x) < eps / 2)
H14:0 < x2 <= M - m
H15:forall z0 : R, Rabs (z0 - x) < x2 -> Rabs (f0 z0 - f0 x) < eps / 2

Rabs (z - x) < x2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x2) < del -> Rabs (f0 z0 - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x2) < zeta -> Rabs (f0 z0 - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m
z:R
H10:Rabs (z - x) < {| pos := x1; cond_pos := H8 |}
exists alp : R, Rabs (z - x) < alp <= x1 /\ E alp
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m
is_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
0 < x1 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x2) < del -> Rabs (f0 z0 - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x2) < zeta -> Rabs (f0 z0 - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m
z:R
H10:Rabs (z - x) < {| pos := x1; cond_pos := H8 |}

exists alp : R, Rabs (z - x) < alp <= x1 /\ E alp
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m
is_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
0 < x1 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x2) < del -> Rabs (f0 z0 - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x2) < zeta -> Rabs (f0 z0 - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m
z:R
H10:Rabs (z - x) < {| pos := x1; cond_pos := H8 |}
H11:exists alp : R, Rabs (z - x) < alp <= x1 /\ E alp

exists alp : R, Rabs (z - x) < alp <= x1 /\ E alp
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x2) < del -> Rabs (f0 z0 - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x2) < zeta -> Rabs (f0 z0 - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m
z:R
H10:Rabs (z - x) < {| pos := x1; cond_pos := H8 |}
H11:~ (exists alp : R, Rabs (z - x) < alp <= x1 /\ E alp)
exists alp : R, Rabs (z - x) < alp <= x1 /\ E alp
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m
is_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
0 < x1 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x2) < del -> Rabs (f0 z0 - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x2) < zeta -> Rabs (f0 z0 - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m
z:R
H10:Rabs (z - x) < {| pos := x1; cond_pos := H8 |}
H11:~ (exists alp : R, Rabs (z - x) < alp <= x1 /\ E alp)

exists alp : R, Rabs (z - x) < alp <= x1 /\ E alp
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m
is_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
0 < x1 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x2) < del -> Rabs (f0 z0 - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x2) < zeta -> Rabs (f0 z0 - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_upper_bound E x1 /\ (forall b : R, is_upper_bound E b -> x1 <= b)
H8:0 < x1
H9:x1 <= M - m
z:R
H10:Rabs (z - x) < {| pos := x1; cond_pos := H8 |}
H11:~ (exists alp : R, Rabs (z - x) < alp <= x1 /\ E alp)
H12:forall n : R, ~ (fun alp : R => Rabs (z - x) < alp <= x1 /\ E alp) n
H13:is_upper_bound E x1
H14:forall b : R, is_upper_bound E b -> x1 <= b

is_upper_bound E (Rabs (z - x)) -> exists alp : R, Rabs (z - x) < alp <= x1 /\ E alp
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x2) < del -> Rabs (f0 z0 - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x2) < zeta -> Rabs (f0 z0 - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_upper_bound E x1 /\ (forall b : R, is_upper_bound E b -> x1 <= b)
H8:0 < x1
H9:x1 <= M - m
z:R
H10:Rabs (z - x) < {| pos := x1; cond_pos := H8 |}
H11:~ (exists alp : R, Rabs (z - x) < alp <= x1 /\ E alp)
H12:forall n : R, ~ (fun alp : R => Rabs (z - x) < alp <= x1 /\ E alp) n
H13:is_upper_bound E x1
H14:forall b : R, is_upper_bound E b -> x1 <= b
is_upper_bound E (Rabs (z - x))
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m
is_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
0 < x1 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x2) < del -> Rabs (f0 z0 - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x2) < zeta -> Rabs (f0 z0 - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_upper_bound E x1 /\ (forall b : R, is_upper_bound E b -> x1 <= b)
H8:0 < x1
H9:x1 <= M - m
z:R
H10:Rabs (z - x) < {| pos := x1; cond_pos := H8 |}
H11:~ (exists alp : R, Rabs (z - x) < alp <= x1 /\ E alp)
H12:forall n : R, ~ (fun alp : R => Rabs (z - x) < alp <= x1 /\ E alp) n
H13:is_upper_bound E x1
H14:forall b : R, is_upper_bound E b -> x1 <= b

is_upper_bound E (Rabs (z - x))
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m
is_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
0 < x1 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x3) < del -> Rabs (f0 z0 - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x3) < zeta -> Rabs (f0 z0 - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x3 : R, D_x no_cond x x3 /\ Rabs (x3 - x) < alp -> Rabs (f0 x3 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x3 : R, D_x no_cond x x3 /\ Rabs (x3 - x) < x0 -> Rabs (f0 x3 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x3 : R, E x3
x1:R
p:is_upper_bound E x1 /\ (forall b : R, is_upper_bound E b -> x1 <= b)
H8:0 < x1
H9:x1 <= M - m
z:R
H10:Rabs (z - x) < {| pos := x1; cond_pos := H8 |}
H11:~ (exists alp : R, Rabs (z - x) < alp <= x1 /\ E alp)
H12:forall n : R, ~ (fun alp : R => Rabs (z - x) < alp <= x1 /\ E alp) n
H13:forall x3 : R, E x3 -> x3 <= x1
H14:forall b : R, is_upper_bound E b -> x1 <= b
x2:R
H15:E x2
H16:x2 <= x1
r:x2 <= Rabs (z - x)

x2 <= Rabs (z - x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x3) < del -> Rabs (f0 z0 - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x3) < zeta -> Rabs (f0 z0 - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x3 : R, D_x no_cond x x3 /\ Rabs (x3 - x) < alp -> Rabs (f0 x3 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x3 : R, D_x no_cond x x3 /\ Rabs (x3 - x) < x0 -> Rabs (f0 x3 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x3 : R, E x3
x1:R
p:is_upper_bound E x1 /\ (forall b : R, is_upper_bound E b -> x1 <= b)
H8:0 < x1
H9:x1 <= M - m
z:R
H10:Rabs (z - x) < {| pos := x1; cond_pos := H8 |}
H11:~ (exists alp : R, Rabs (z - x) < alp <= x1 /\ E alp)
H12:forall n0 : R, ~ (fun alp : R => Rabs (z - x) < alp <= x1 /\ E alp) n0
H13:forall x3 : R, E x3 -> x3 <= x1
H14:forall b : R, is_upper_bound E b -> x1 <= b
x2:R
H15:E x2
H16:x2 <= x1
n:~ x2 <= Rabs (z - x)
x2 <= Rabs (z - x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m
is_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
0 < x1 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x3) < del -> Rabs (f0 z0 - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x3) < zeta -> Rabs (f0 z0 - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x3 : R, D_x no_cond x x3 /\ Rabs (x3 - x) < alp -> Rabs (f0 x3 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x3 : R, D_x no_cond x x3 /\ Rabs (x3 - x) < x0 -> Rabs (f0 x3 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x3 : R, E x3
x1:R
p:is_upper_bound E x1 /\ (forall b : R, is_upper_bound E b -> x1 <= b)
H8:0 < x1
H9:x1 <= M - m
z:R
H10:Rabs (z - x) < {| pos := x1; cond_pos := H8 |}
H11:~ (exists alp : R, Rabs (z - x) < alp <= x1 /\ E alp)
H12:forall n0 : R, ~ (fun alp : R => Rabs (z - x) < alp <= x1 /\ E alp) n0
H13:forall x3 : R, E x3 -> x3 <= x1
H14:forall b : R, is_upper_bound E b -> x1 <= b
x2:R
H15:E x2
H16:x2 <= x1
n:~ x2 <= Rabs (z - x)

x2 <= Rabs (z - x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m
is_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
0 < x1 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m

is_lub E {| pos := x1; cond_pos := H8 |} /\ disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
0 < x1 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m

is_lub E {| pos := x1; cond_pos := H8 |}
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m
disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
0 < x1 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
H8:0 < x1
H9:x1 <= M - m

disc x {| pos := {| pos := x1; cond_pos := H8 |} / 2; cond_pos := H1 {| pos := x1; cond_pos := H8 |} |} x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1
0 < x1 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < alp -> Rabs (f0 x2 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x2 : R, D_x no_cond x x2 /\ Rabs (x2 - x) < x0 -> Rabs (f0 x2 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x2 : R, E x2
x1:R
p:is_lub E x1

0 < x1 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x3 : R, D_x no_cond x x3 /\ Rabs (x3 - x) < alp -> Rabs (f0 x3 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x3 : R, D_x no_cond x x3 /\ Rabs (x3 - x) < x0 -> Rabs (f0 x3 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x3 : R, E x3
x1:R
p:is_upper_bound E x1 /\ (forall b : R, is_upper_bound E b -> x1 <= b)
x2:R
H8:0 < x2 <= M - m /\ (forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2)
H9:0 < x2 <= M - m
H10:0 < x2
H11:forall x3 : R, E x3 -> x3 <= x1
H12:forall b : R, (forall x3 : R, E x3 -> x3 <= b) -> x1 <= b

0 < x1
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x3 : R, D_x no_cond x x3 /\ Rabs (x3 - x) < alp -> Rabs (f0 x3 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x3 : R, D_x no_cond x x3 /\ Rabs (x3 - x) < x0 -> Rabs (f0 x3 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x3 : R, E x3
x1:R
p:is_upper_bound E x1 /\ (forall b : R, is_upper_bound E b -> x1 <= b)
x2:R
H8:0 < x2 <= M - m /\ (forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2)
H9:0 < x2 <= M - m
H10:0 < x2
H11:forall x3 : R, E x3 -> x3 <= x1
H12:forall b : R, (forall x3 : R, E x3 -> x3 <= b) -> x1 <= b
x1 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
H4:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x3 : R, D_x no_cond x x3 /\ Rabs (x3 - x) < alp -> Rabs (f0 x3 - f0 x) < eps0)
x0:R
H5:x0 > 0 /\ (forall x3 : R, D_x no_cond x x3 /\ Rabs (x3 - x) < x0 -> Rabs (f0 x3 - f0 x) < eps / 2)
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H6:bound E
H7:exists x3 : R, E x3
x1:R
p:is_upper_bound E x1 /\ (forall b : R, is_upper_bound E b -> x1 <= b)
x2:R
H8:0 < x2 <= M - m /\ (forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2)
H9:0 < x2 <= M - m
H10:0 < x2
H11:forall x3 : R, E x3 -> x3 <= x1
H12:forall b : R, (forall x3 : R, E x3 -> x3 <= b) -> x1 <= b

x1 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family

family_open_set f'
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x

open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x = x0

exists delta : posreal, included (disc x0 delta) (fun y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y))
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
exists delta : posreal, included (disc x0 delta) (fun y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y))
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x = x0
x2:R
H7:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2

X x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x = x0
x2:R
H7:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2
exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
exists delta : posreal, included (disc x0 delta) (fun y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y))
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x = x0
x2:R
H7:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2

exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
exists delta : posreal, included (disc x0 delta) (fun y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y))
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x = x0
x2:R
H7:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2

forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x = x0
x2:R
H7:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2
is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
exists delta : posreal, included (disc x0 delta) (fun y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y))
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x = x0
x2:R
H7:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2

is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
exists delta : posreal, included (disc x0 delta) (fun y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y))
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x = x0
x2:R
H7:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2

is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x = x0
x2:R
H7:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2
disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
exists delta : posreal, included (disc x0 delta) (fun y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y))
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x = x0
x2:R
H7:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2

disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
exists delta : posreal, included (disc x0 delta) (fun y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y))
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0

exists delta : posreal, included (disc x0 delta) (fun y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y))
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
d:=x1 / 2 - Rabs (x0 - x):R

0 < d
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
d:=x1 / 2 - Rabs (x0 - x):R
H7:0 < d
exists delta : posreal, included (disc x0 delta) (fun y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y))
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
d:=x1 / 2 - Rabs (x0 - x):R
H7:0 < d

exists delta : posreal, included (disc x0 delta) (fun y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y))
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
d:=x1 / 2 - Rabs (x0 - x):R
H7:0 < d
x2:R
H8:disc x0 {| pos := d; cond_pos := H7 |} x2

X x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
d:=x1 / 2 - Rabs (x0 - x):R
H7:0 < d
x2:R
H8:disc x0 {| pos := d; cond_pos := H7 |} x2
exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
d:=x1 / 2 - Rabs (x0 - x):R
H7:0 < d
x2:R
H8:disc x0 {| pos := d; cond_pos := H7 |} x2

exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
d:=x1 / 2 - Rabs (x0 - x):R
H7:0 < d
x2:R
H8:disc x0 {| pos := d; cond_pos := H7 |} x2

forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
d:=x1 / 2 - Rabs (x0 - x):R
H7:0 < d
x2:R
H8:disc x0 {| pos := d; cond_pos := H7 |} x2
is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
d:=x1 / 2 - Rabs (x0 - x):R
H7:0 < d
x2:R
H8:disc x0 {| pos := d; cond_pos := H7 |} x2

is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
d:=x1 / 2 - Rabs (x0 - x):R
H7:0 < d
x2:R
H8:disc x0 {| pos := d; cond_pos := H7 |} x2
H9:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1
H10:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0

is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
d:=x1 / 2 - Rabs (x0 - x):R
H7:0 < d
x2:R
H8:disc x0 {| pos := d; cond_pos := H7 |} x2
H9:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1
H10:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
d:=x1 / 2 - Rabs (x0 - x):R
H7:0 < d
x2:R
H8:disc x0 {| pos := d; cond_pos := H7 |} x2
H9:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1
H10:disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0

disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
d:=x1 / 2 - Rabs (x0 - x):R
H7:0 < d
x2:R
H8:Rabs (x2 - x0) < d
H9:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1
H10:Rabs (x0 - x) < x1 / 2

Rabs (x2 - x) <= Rabs (x2 - x0) + Rabs (x0 - x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
d:=x1 / 2 - Rabs (x0 - x):R
H7:0 < d
x2:R
H8:Rabs (x2 - x0) < d
H9:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1
H10:Rabs (x0 - x) < x1 / 2
Rabs (x2 - x0) + Rabs (x0 - x) < x1 / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
d:=x1 / 2 - Rabs (x0 - x):R
H7:0 < d
x2:R
H8:Rabs (x2 - x0) < d
H9:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1
H10:Rabs (x0 - x) < x1 / 2

Rabs (x2 - x0) + Rabs (x0 - x) < x1 / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:X x
x0:R
x1:posreal
H4:forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2
H5:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1 /\ disc x {| pos := x1 / 2; cond_pos := H1 x1 |} x0
H6:x <> x0
d:=x1 / 2 - Rabs (x0 - x):R
H7:0 < d
x2:R
H8:Rabs (x2 - x0) < d
H9:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x1
H10:Rabs (x0 - x) < x1 / 2

Rabs (x2 - x0) + Rabs (x0 - x) < d + Rabs (x0 - x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x

open_set (g x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x

open_set (fun _ : R => False)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
(fun _ : R => False) =_D g x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x

(fun _ : R => False) =_D g x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x

forall x0 : R, False -> g x x0
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x
forall x0 : R, g x x0 -> False
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
x:R
H3:~ X x

forall x0 : R, g x x0 -> False
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'

exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l

(forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)) -> exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R

0 < D / 2 -> exists delta : posreal, forall x y : R, X x -> X y -> Rabs (x - y) < delta -> Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi

In xi l
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l

Rabs (f0 x - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i

Rabs (f0 x - f0 y) <= Rabs (f0 x - f0 xi) + Rabs (f0 xi - f0 y)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
Rabs (f0 x - f0 xi) + Rabs (f0 xi - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i

Rabs (f0 x - f0 xi) + Rabs (f0 xi - f0 y) < eps
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i

Rabs (f0 x - f0 xi) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
Rabs (f0 xi - f0 y) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2

Rabs (x - xi) < pos_Rl l' i / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2
pos_Rl l' i / 2 < pos_Rl l' i
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
Rabs (f0 xi - f0 y) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2

g xi x
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2
pos_Rl l' i / 2 < pos_Rl l' i
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
Rabs (f0 xi - f0 y) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2

pos_Rl l' i / 2 < pos_Rl l' i
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
Rabs (f0 xi - f0 y) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2

0 < 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2
2 * (pos_Rl l' i * / 2) < 2 * pos_Rl l' i
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
Rabs (f0 xi - f0 y) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2

2 * (pos_Rl l' i * / 2) < 2 * pos_Rl l' i
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
Rabs (f0 xi - f0 y) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2

pos_Rl l' i * 1 < 2 * pos_Rl l' i
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2
2 <> 0
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
Rabs (f0 xi - f0 y) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2

2 <> 0
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
Rabs (f0 xi - f0 y) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i

Rabs (f0 xi - f0 y) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2
H22:g xi x
H23:DF xi
H24:Rabs (x - xi) < pos_Rl l' i / 2

Rabs (y - xi) <= Rabs (y - x) + Rabs (x - xi)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2
H22:g xi x
H23:DF xi
H24:Rabs (x - xi) < pos_Rl l' i / 2
Rabs (y - x) + Rabs (x - xi) < pos_Rl l' i
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2
H22:g xi x
H23:DF xi
H24:Rabs (x - xi) < pos_Rl l' i / 2

Rabs (y - x) + Rabs (x - xi) < pos_Rl l' i
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2
H22:g xi x
H23:DF xi
H24:Rabs (x - xi) < pos_Rl l' i / 2

Rabs (y - x) < pos_Rl l' i / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2
H22:g xi x
H23:DF xi
H24:Rabs (x - xi) < pos_Rl l' i / 2
Rabs (x - xi) < pos_Rl l' i / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2
H22:g xi x
H23:DF xi
H24:Rabs (x - xi) < pos_Rl l' i / 2

Rabs (y - x) < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2
H22:g xi x
H23:DF xi
H24:Rabs (x - xi) < pos_Rl l' i / 2
D / 2 <= pos_Rl l' i / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2
H22:g xi x
H23:DF xi
H24:Rabs (x - xi) < pos_Rl l' i / 2
Rabs (x - xi) < pos_Rl l' i / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2
H22:g xi x
H23:DF xi
H24:Rabs (x - xi) < pos_Rl l' i / 2

D / 2 <= pos_Rl l' i / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2
H22:g xi x
H23:DF xi
H24:Rabs (x - xi) < pos_Rl l' i / 2
Rabs (x - xi) < pos_Rl l' i / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2
H22:g xi x
H23:DF xi
H24:Rabs (x - xi) < pos_Rl l' i / 2

0 <= / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2
H22:g xi x
H23:DF xi
H24:Rabs (x - xi) < pos_Rl l' i / 2
D <= pos_Rl l' i
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2
H22:g xi x
H23:DF xi
H24:Rabs (x - xi) < pos_Rl l' i / 2
Rabs (x - xi) < pos_Rl l' i / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2
H22:g xi x
H23:DF xi
H24:Rabs (x - xi) < pos_Rl l' i / 2

D <= pos_Rl l' i
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2
H22:g xi x
H23:DF xi
H24:Rabs (x - xi) < pos_Rl l' i / 2
Rabs (x - xi) < pos_Rl l' i / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y0 : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y0):R -> R -> Prop
H2:forall x0 : R, (exists y0 : R, g x0 y0) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y0 : R, g y0 x0 /\ DF y0
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
H6:forall x0 : R, In x0 l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ included (g x0) (fun z : R => Rabs (z - x0) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i0 : nat, (i0 < Rlength l)%nat -> 0 < pos_Rl l' i0 /\ (forall z : R, Rabs (z - pos_Rl l i0) < pos_Rl l' i0 -> Rabs (f0 z - f0 (pos_Rl l i0)) < eps / 2) /\ included (g (pos_Rl l i0)) (fun z : R => Rabs (z - pos_Rl l i0) < pos_Rl l' i0 / 2)
D:=MinRlist l':R
H9:0 < D / 2
x, y:R
H10:X x
H11:X y
H12:Rabs (x - y) < {| pos := D / 2; cond_pos := H9 |}
xi:R
H13:g xi x /\ DF xi
H14:In xi l
H15:In xi l -> exists i0 : nat, (i0 < Rlength l)%nat /\ xi = pos_Rl l i0
i:nat
H16:(i < Rlength l)%nat /\ xi = pos_Rl l i
H17:(i < Rlength l)%nat
H18:xi = pos_Rl l i
H19:0 < pos_Rl l' i
H20:forall z : R, Rabs (z - xi) < pos_Rl l' i -> Rabs (f0 z - f0 xi) < eps / 2
H21:forall x0 : R, g xi x0 -> Rabs (x0 - xi) < pos_Rl l' i / 2
H22:g xi x
H23:DF xi
H24:Rabs (x - xi) < pos_Rl l' i / 2

Rabs (x - xi) < pos_Rl l' i / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R
0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D0 : R -> Prop, covering_finite X (subfamily f1 D0)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
H6:forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
l':Rlist
H7:Rlength l = Rlength l'
H8:forall i : nat, (i < Rlength l)%nat -> 0 < pos_Rl l' i /\ (forall z : R, Rabs (z - pos_Rl l i) < pos_Rl l' i -> Rabs (f0 z - f0 (pos_Rl l i)) < eps / 2) /\ included (g (pos_Rl l i)) (fun z : R => Rabs (z - pos_Rl l i) < pos_Rl l' i / 2)
D:=MinRlist l':R

0 < D / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l
forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x : R, X x -> continuity_pt f0 x
m, M:R
X_enc:forall x : R, X x -> m <= x <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x y : R => X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x : R, (exists y : R, g x y) -> X x
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x : R, X x -> exists y : R, g y x /\ DF y
l:Rlist
H5:forall x : R, X x /\ DF x <-> In x l

forall x : R, In x l -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y : R, g y x0 /\ DF y
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop

bound E
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y : R, g y x0 /\ DF y
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y : R, g y x0 /\ DF y
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E

exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y : R, g y x0 /\ DF y
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E

exists x0 : R, E x0
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y : R, g y x0 /\ DF y
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x0 : R, E x0
exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H13:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H12:x0 > 0
H14:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2

0 < Rmin x0 (M - m) <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H13:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H12:x0 > 0
H14:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2
forall z : R, Rabs (z - x) < Rmin x0 (M - m) -> Rabs (f0 z - f0 x) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y : R, g y x0 /\ DF y
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x0 : R, E x0
exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H13:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H12:x0 > 0
H14:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2

forall z : R, Rabs (z - x) < Rmin x0 (M - m) -> Rabs (f0 z - f0 x) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y : R, g y x0 /\ DF y
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x0 : R, E x0
exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x1) < del -> Rabs (f0 z0 - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x1) < zeta -> Rabs (f0 z0 - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H11:bound E
H13:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H12:x0 > 0
H14:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2
z:R
H15:Rabs (z - x) < Rmin x0 (M - m)
H16:x = z

Rabs (f0 z - f0 x) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x1) < del -> Rabs (f0 z0 - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x1) < zeta -> Rabs (f0 z0 - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H11:bound E
H13:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H12:x0 > 0
H14:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2
z:R
H15:Rabs (z - x) < Rmin x0 (M - m)
H16:x <> z
Rabs (f0 z - f0 x) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y : R, g y x0 /\ DF y
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x0 : R, E x0
exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x1) < del -> Rabs (f0 z0 - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x1) < zeta -> Rabs (f0 z0 - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H11:bound E
H13:forall eps0 : R, eps0 > 0 -> exists alp : R, alp > 0 /\ (forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < alp -> Rabs (f0 x1 - f0 x) < eps0)
x0:R
H12:x0 > 0
H14:forall x1 : R, D_x no_cond x x1 /\ Rabs (x1 - x) < x0 -> Rabs (f0 x1 - f0 x) < eps / 2
z:R
H15:Rabs (z - x) < Rmin x0 (M - m)
H16:x <> z

Rabs (f0 z - f0 x) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y : R, g y x0 /\ DF y
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x0 : R, E x0
exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x0 : R, X x0 -> continuity_pt f0 x0
m, M:R
X_enc:forall x0 : R, X x0 -> m <= x0 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x0 y : R => X x0 /\ (exists del : posreal, (forall z : R, Rabs (z - x0) < del -> Rabs (f0 z - f0 x0) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x0) < zeta -> Rabs (f0 z - f0 x0) < eps / 2)) del /\ disc x0 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x0 : R, (exists y : R, g x0 y) -> X x0
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x0 : R, X x0 -> exists y : R, g y x0 /\ DF y
l:Rlist
H5:forall x0 : R, X x0 /\ DF x0 <-> In x0 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x0 : R, E x0

exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0

exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0

0 < x0 <= M - m -> exists del : R, 0 < del /\ (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < del / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
0 < x0 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m

0 < x0
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m
(forall z : R, Rabs (z - x) < x0 -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < x0 / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
0 < x0 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m

(forall z : R, Rabs (z - x) < x0 -> Rabs (f0 z - f0 x) < eps / 2) /\ included (g x) (fun z : R => Rabs (z - x) < x0 / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
0 < x0 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m

forall z : R, Rabs (z - x) < x0 -> Rabs (f0 z - f0 x) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m
included (g x) (fun z : R => Rabs (z - x) < x0 / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
0 < x0 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x1) < del -> Rabs (f0 z0 - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x1) < zeta -> Rabs (f0 z0 - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m
z:R
H15:Rabs (z - x) < x0

(exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp) -> Rabs (f0 z - f0 x) < eps / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x1) < del -> Rabs (f0 z0 - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x1) < zeta -> Rabs (f0 z0 - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m
z:R
H15:Rabs (z - x) < x0
exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m
included (g x) (fun z : R => Rabs (z - x) < x0 / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
0 < x0 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x1) < del -> Rabs (f0 z0 - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x1) < zeta -> Rabs (f0 z0 - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m
z:R
H15:Rabs (z - x) < x0

exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m
included (g x) (fun z : R => Rabs (z - x) < x0 / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
0 < x0 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x1) < del -> Rabs (f0 z0 - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x1) < zeta -> Rabs (f0 z0 - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m
z:R
H15:Rabs (z - x) < x0
H16:exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp

exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x1) < del -> Rabs (f0 z0 - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x1) < zeta -> Rabs (f0 z0 - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m
z:R
H15:Rabs (z - x) < x0
H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)
exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m
included (g x) (fun z : R => Rabs (z - x) < x0 / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
0 < x0 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x1) < del -> Rabs (f0 z0 - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x1) < zeta -> Rabs (f0 z0 - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m
z:R
H15:Rabs (z - x) < x0
H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)

exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m
included (g x) (fun z : R => Rabs (z - x) < x0 / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
0 < x0 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x1) < del -> Rabs (f0 z0 - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x1) < zeta -> Rabs (f0 z0 - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)
H13:0 < x0
H14:x0 <= M - m
z:R
H15:Rabs (z - x) < x0
H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)
H17:forall n : R, ~ (fun alp : R => Rabs (z - x) < alp <= x0 /\ E alp) n
H18:is_upper_bound E x0
H19:forall b : R, is_upper_bound E b -> x0 <= b

is_upper_bound E (Rabs (z - x)) -> exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x1) < del -> Rabs (f0 z0 - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x1) < zeta -> Rabs (f0 z0 - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)
H13:0 < x0
H14:x0 <= M - m
z:R
H15:Rabs (z - x) < x0
H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)
H17:forall n : R, ~ (fun alp : R => Rabs (z - x) < alp <= x0 /\ E alp) n
H18:is_upper_bound E x0
H19:forall b : R, is_upper_bound E b -> x0 <= b
is_upper_bound E (Rabs (z - x))
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m
included (g x) (fun z : R => Rabs (z - x) < x0 / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
0 < x0 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x1) < del -> Rabs (f0 z0 - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x1) < zeta -> Rabs (f0 z0 - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)
H13:0 < x0
H14:x0 <= M - m
z:R
H15:Rabs (z - x) < x0
H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)
H17:forall n : R, ~ (fun alp : R => Rabs (z - x) < alp <= x0 /\ E alp) n
H18:is_upper_bound E x0
H19:forall b : R, is_upper_bound E b -> x0 <= b

is_upper_bound E (Rabs (z - x))
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m
included (g x) (fun z : R => Rabs (z - x) < x0 / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
0 < x0 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x2) < del -> Rabs (f0 z0 - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x2) < zeta -> Rabs (f0 z0 - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x2 : R, X x2 -> exists y : R, g y x2 /\ DF y
l:Rlist
H5:forall x2 : R, X x2 /\ DF x2 <-> In x2 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x2 : R, E x2
x0:R
p:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)
H13:0 < x0
H14:x0 <= M - m
z:R
H15:Rabs (z - x) < x0
H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)
H17:forall n : R, ~ (fun alp : R => Rabs (z - x) < alp <= x0 /\ E alp) n
H18:forall x2 : R, E x2 -> x2 <= x0
H19:forall b : R, is_upper_bound E b -> x0 <= b
x1:R
H20:E x1
H21:x1 <= x0
r:x1 <= Rabs (z - x)

x1 <= Rabs (z - x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x2) < del -> Rabs (f0 z0 - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x2) < zeta -> Rabs (f0 z0 - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x2 : R, X x2 -> exists y : R, g y x2 /\ DF y
l:Rlist
H5:forall x2 : R, X x2 /\ DF x2 <-> In x2 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x2 : R, E x2
x0:R
p:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)
H13:0 < x0
H14:x0 <= M - m
z:R
H15:Rabs (z - x) < x0
H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)
H17:forall n0 : R, ~ (fun alp : R => Rabs (z - x) < alp <= x0 /\ E alp) n0
H18:forall x2 : R, E x2 -> x2 <= x0
H19:forall b : R, is_upper_bound E b -> x0 <= b
x1:R
H20:E x1
H21:x1 <= x0
n:~ x1 <= Rabs (z - x)
x1 <= Rabs (z - x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m
included (g x) (fun z : R => Rabs (z - x) < x0 / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
0 < x0 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x2) < del -> Rabs (f0 z0 - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x2) < zeta -> Rabs (f0 z0 - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x2 : R, X x2 -> exists y : R, g y x2 /\ DF y
l:Rlist
H5:forall x2 : R, X x2 /\ DF x2 <-> In x2 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x2 : R, E x2
x0:R
p:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)
H13:0 < x0
H14:x0 <= M - m
z:R
H15:Rabs (z - x) < x0
H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)
H17:forall n0 : R, ~ (fun alp : R => Rabs (z - x) < alp <= x0 /\ E alp) n0
H18:forall x2 : R, E x2 -> x2 <= x0
H19:forall b : R, is_upper_bound E b -> x0 <= b
x1:R
H20:E x1
H21:x1 <= x0
n:~ x1 <= Rabs (z - x)

x1 <= Rabs (z - x)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m
included (g x) (fun z : R => Rabs (z - x) < x0 / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
0 < x0 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x2) < del -> Rabs (f0 z0 - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x2) < zeta -> Rabs (f0 z0 - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x2 : R, X x2 -> exists y : R, g y x2 /\ DF y
l:Rlist
H5:forall x2 : R, X x2 /\ DF x2 <-> In x2 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x2 : R, E x2
x0:R
p:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)
H13:0 < x0
H14:x0 <= M - m
z:R
H15:Rabs (z - x) < x0
H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)
H17:forall n0 : R, ~ (fun alp : R => Rabs (z - x) < alp <= x0 /\ E alp) n0
H18:forall x2 : R, E x2 -> x2 <= x0
H19:forall b : R, is_upper_bound E b -> x0 <= b
x1:R
H20:E x1
H21:x1 <= x0
n:~ x1 <= Rabs (z - x)

Rabs (z - x) < x1 <= x0
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x2) < del -> Rabs (f0 z0 - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x2) < zeta -> Rabs (f0 z0 - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x2 : R, X x2 -> exists y : R, g y x2 /\ DF y
l:Rlist
H5:forall x2 : R, X x2 /\ DF x2 <-> In x2 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x2 : R, E x2
x0:R
p:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)
H13:0 < x0
H14:x0 <= M - m
z:R
H15:Rabs (z - x) < x0
H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)
H17:forall n0 : R, ~ (fun alp : R => Rabs (z - x) < alp <= x0 /\ E alp) n0
H18:forall x2 : R, E x2 -> x2 <= x0
H19:forall b : R, is_upper_bound E b -> x0 <= b
x1:R
H20:E x1
H21:x1 <= x0
n:~ x1 <= Rabs (z - x)
E x1
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m
included (g x) (fun z : R => Rabs (z - x) < x0 / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
0 < x0 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z0 : R, Rabs (z0 - x2) < del -> Rabs (f0 z0 - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x2) < zeta -> Rabs (f0 z0 - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x2 : R, X x2 -> exists y : R, g y x2 /\ DF y
l:Rlist
H5:forall x2 : R, X x2 /\ DF x2 <-> In x2 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z0 : R, Rabs (z0 - x) < zeta -> Rabs (f0 z0 - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x2 : R, E x2
x0:R
p:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)
H13:0 < x0
H14:x0 <= M - m
z:R
H15:Rabs (z - x) < x0
H16:~ (exists alp : R, Rabs (z - x) < alp <= x0 /\ E alp)
H17:forall n0 : R, ~ (fun alp : R => Rabs (z - x) < alp <= x0 /\ E alp) n0
H18:forall x2 : R, E x2 -> x2 <= x0
H19:forall b : R, is_upper_bound E b -> x0 <= b
x1:R
H20:E x1
H21:x1 <= x0
n:~ x1 <= Rabs (z - x)

E x1
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m
included (g x) (fun z : R => Rabs (z - x) < x0 / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
0 < x0 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m

included (g x) (fun z : R => Rabs (z - x) < x0 / 2)
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
0 < x0 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x3 : R, X x3 -> exists y : R, g y x3 /\ DF y
l:Rlist
H5:forall x3 : R, X x3 /\ DF x3 <-> In x3 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x3 : R, E x3
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m
x1:R
H15:X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x1)
H16:X x
H17:exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x1
x2:posreal
H18:(forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x2 /\ disc x {| pos := x2 / 2; cond_pos := H1 x2 |} x1
H19:forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2
H21:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x2
H22:disc x {| pos := x2 / 2; cond_pos := H1 x2 |} x1

x0 = x2 -> Rabs (x1 - x) < x0 / 2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x3 : R, X x3 -> exists y : R, g y x3 /\ DF y
l:Rlist
H5:forall x3 : R, X x3 /\ DF x3 <-> In x3 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x3 : R, E x3
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m
x1:R
H15:X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x1)
H16:X x
H17:exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x1
x2:posreal
H18:(forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x2 /\ disc x {| pos := x2 / 2; cond_pos := H1 x2 |} x1
H19:forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2
H21:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x2
H22:disc x {| pos := x2 / 2; cond_pos := H1 x2 |} x1
x0 = x2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
0 < x0 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x3 : R, X x3 -> exists y : R, g y x3 /\ DF y
l:Rlist
H5:forall x3 : R, X x3 /\ DF x3 <-> In x3 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x3 : R, E x3
x0:R
p:is_lub E x0
H13:0 < x0
H14:x0 <= M - m
x1:R
H15:X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x1)
H16:X x
H17:exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x1
x2:posreal
H18:(forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x2 /\ disc x {| pos := x2 / 2; cond_pos := H1 x2 |} x1
H19:forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2
H21:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x2
H22:disc x {| pos := x2 / 2; cond_pos := H1 x2 |} x1

x0 = x2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
0 < x0 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x3 : R, X x3 -> exists y : R, g y x3 /\ DF y
l:Rlist
H5:forall x3 : R, X x3 /\ DF x3 <-> In x3 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x3 : R, E x3
x0:R
p:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x0
H13:0 < x0
H14:x0 <= M - m
x1:R
H15:X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x1)
H16:X x
H17:exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x1
x2:posreal
H18:(forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x2 /\ disc x {| pos := x2 / 2; cond_pos := H1 x2 |} x1
H19:forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2
H21:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x2
H22:disc x {| pos := x2 / 2; cond_pos := H1 x2 |} x1

is_lub ?E x0
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x3 : R, X x3 -> exists y : R, g y x3 /\ DF y
l:Rlist
H5:forall x3 : R, X x3 /\ DF x3 <-> In x3 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x3 : R, E x3
x0:R
p:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x0
H13:0 < x0
H14:x0 <= M - m
x1:R
H15:X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x1)
H16:X x
H17:exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x1
x2:posreal
H18:(forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x2 /\ disc x {| pos := x2 / 2; cond_pos := H1 x2 |} x1
H19:forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2
H21:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x2
H22:disc x {| pos := x2 / 2; cond_pos := H1 x2 |} x1
is_lub ?E x2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
0 < x0 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x3 : R, X x3 -> continuity_pt f0 x3
m, M:R
X_enc:forall x3 : R, X x3 -> m <= x3 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x3 y : R => X x3 /\ (exists del : posreal, (forall z : R, Rabs (z - x3) < del -> Rabs (f0 z - f0 x3) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x3) < zeta -> Rabs (f0 z - f0 x3) < eps / 2)) del /\ disc x3 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x3 : R, (exists y : R, g x3 y) -> X x3
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x3 : R, X x3 -> exists y : R, g y x3 /\ DF y
l:Rlist
H5:forall x3 : R, X x3 /\ DF x3 <-> In x3 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x3 : R, E x3
x0:R
p:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x0
H13:0 < x0
H14:x0 <= M - m
x1:R
H15:X x /\ (exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x1)
H16:X x
H17:exists del : posreal, (forall z : R, Rabs (z - x) < del -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) del /\ disc x {| pos := del / 2; cond_pos := H1 del |} x1
x2:posreal
H18:(forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x2 /\ disc x {| pos := x2 / 2; cond_pos := H1 x2 |} x1
H19:forall z : R, Rabs (z - x) < x2 -> Rabs (f0 z - f0 x) < eps / 2
H21:is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x2
H22:disc x {| pos := x2 / 2; cond_pos := H1 x2 |} x1

is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2)) x2
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0
0 < x0 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x1 : R, X x1 -> continuity_pt f0 x1
m, M:R
X_enc:forall x1 : R, X x1 -> m <= x1 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x1 y : R => X x1 /\ (exists del : posreal, (forall z : R, Rabs (z - x1) < del -> Rabs (f0 z - f0 x1) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x1) < zeta -> Rabs (f0 z - f0 x1) < eps / 2)) del /\ disc x1 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x1 : R, (exists y : R, g x1 y) -> X x1
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x1 : R, X x1 -> exists y : R, g y x1 /\ DF y
l:Rlist
H5:forall x1 : R, X x1 /\ DF x1 <-> In x1 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x1 : R, E x1
x0:R
p:is_lub E x0

0 < x0 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x2 : R, X x2 -> exists y : R, g y x2 /\ DF y
l:Rlist
H5:forall x2 : R, X x2 /\ DF x2 <-> In x2 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x2 : R, E x2
x0:R
p:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)
x1:R
H13:0 < x1 <= M - m /\ (forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2)
H14:0 < x1 <= M - m
H15:0 < x1
H16:forall x2 : R, E x2 -> x2 <= x0
H17:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x0 <= b

0 < x0
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x2 : R, X x2 -> exists y : R, g y x2 /\ DF y
l:Rlist
H5:forall x2 : R, X x2 /\ DF x2 <-> In x2 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x2 : R, E x2
x0:R
p:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)
x1:R
H13:0 < x1 <= M - m /\ (forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2)
H14:0 < x1 <= M - m
H15:0 < x1
H16:forall x2 : R, E x2 -> x2 <= x0
H17:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x0 <= b
x0 <= M - m
f0:R -> R
X:R -> Prop
H0:forall f1 : family, covering_open_set X f1 -> exists D : R -> Prop, covering_finite X (subfamily f1 D)
H:forall x2 : R, X x2 -> continuity_pt f0 x2
m, M:R
X_enc:forall x2 : R, X x2 -> m <= x2 <= M
Hyp:m < M
eps:posreal
H1:forall t : posreal, 0 < t / 2
g:=fun x2 y : R => X x2 /\ (exists del : posreal, (forall z : R, Rabs (z - x2) < del -> Rabs (f0 z - f0 x2) < eps / 2) /\ is_lub (fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x2) < zeta -> Rabs (f0 z - f0 x2) < eps / 2)) del /\ disc x2 {| pos := del / 2; cond_pos := H1 del |} y):R -> R -> Prop
H2:forall x2 : R, (exists y : R, g x2 y) -> X x2
f':={| ind := X; f := g; cond_fam := H2 |}:family
H3:covering_open_set X f'
DF:R -> Prop
H4:forall x2 : R, X x2 -> exists y : R, g y x2 /\ DF y
l:Rlist
H5:forall x2 : R, X x2 /\ DF x2 <-> In x2 l
x:R
H6:In x l
H7:X x /\ DF x -> In x l
H8:In x l -> X x /\ DF x
H9:X x
H10:DF x
E:=fun zeta : R => 0 < zeta <= M - m /\ (forall z : R, Rabs (z - x) < zeta -> Rabs (f0 z - f0 x) < eps / 2):R -> Prop
H11:bound E
H12:exists x2 : R, E x2
x0:R
p:is_upper_bound E x0 /\ (forall b : R, is_upper_bound E b -> x0 <= b)
x1:R
H13:0 < x1 <= M - m /\ (forall z : R, Rabs (z - x) < x1 -> Rabs (f0 z - f0 x) < eps / 2)
H14:0 < x1 <= M - m
H15:0 < x1
H16:forall x2 : R, E x2 -> x2 <= x0
H17:forall b : R, (forall x2 : R, E x2 -> x2 <= b) -> x0 <= b

x0 <= M - m
apply H17; intros; unfold E in H18; elim H18; intros; elim H19; intros; assumption. Qed.